thys/Positions.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 06 Jul 2017 16:05:33 +0100
changeset 264 e2828c4a1e23
parent 263 00c9a95d492e
child 265 d36be1e356c0
permissions -rw-r--r--
updated
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
   
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
theory Positions
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
  imports "Lexer" 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
     6
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
     7
section {* Position in values *}
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
     8
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
fun 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
  at :: "val \<Rightarrow> nat list \<Rightarrow> val"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
where
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  "at v [] = v"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
| "at (Left v) (0#ps)= at v ps"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
| "at (Right v) (Suc 0#ps)= at v ps"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
| "at (Seq v1 v2) (0#ps)= at v1 ps"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
| "at (Seq v1 v2) (Suc 0#ps)= at v2 ps"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
| "at (Stars vs) (n#ps)= at (nth vs n) ps"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
fun Pos :: "val \<Rightarrow> (nat list) set"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
where
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
  "Pos (Void) = {[]}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
| "Pos (Char c) = {[]}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
| "Pos (Left v) = {[]} \<union> {0#ps | ps. ps \<in> Pos v}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
| "Pos (Right v) = {[]} \<union> {1#ps | ps. ps \<in> Pos v}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
| "Pos (Seq v1 v2) = {[]} \<union> {0#ps | ps. ps \<in> Pos v1} \<union> {1#ps | ps. ps \<in> Pos v2}" 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
| "Pos (Stars []) = {[]}"
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    27
| "Pos (Stars (v#vs)) = {[]} \<union> {0#ps | ps. ps \<in> Pos v} \<union> {Suc n#ps | n ps. n#ps \<in> Pos (Stars vs)}"
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
253
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    29
lemma Pos_stars:
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    30
  "Pos (Stars vs) = {[]} \<union> (\<Union>n < length vs. {n#ps | ps. ps \<in> Pos (vs ! n)})"
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    31
apply(induct vs)
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    32
apply(simp) 
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    33
apply(simp)
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    34
apply(simp add: insert_ident)
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    35
apply(rule subset_antisym)
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    36
apply(auto)
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    37
apply(auto)[1]
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    38
using less_Suc_eq_0_disj by auto
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    39
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
lemma Pos_empty:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
  shows "[] \<in> Pos v"
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    42
by (induct v rule: Pos.induct)(auto)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
fun intlen :: "'a list \<Rightarrow> int"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
where
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
  "intlen [] = 0"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
| "intlen (x#xs) = 1 + intlen xs"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
255
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
    49
lemma intlen_bigger:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  shows "0 \<le> intlen xs"
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    51
by (induct xs)(auto)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
lemma intlen_append:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
  shows "intlen (xs @ ys) = intlen xs + intlen ys"
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    55
by (induct xs arbitrary: ys) (auto)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
lemma intlen_length:
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
    58
  shows "intlen xs < intlen ys \<longleftrightarrow> length xs < length ys"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
    59
apply(induct xs arbitrary: ys)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    60
apply (auto simp add: intlen_bigger not_less)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    61
apply (metis intlen.elims intlen_bigger le_imp_0_less)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
    62
apply (smt Suc_lessI intlen.simps(2) length_Suc_conv nat_neq_iff)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
by (smt Suc_lessE intlen.simps(2) length_Suc_conv)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
255
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
    65
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
definition pflat_len :: "val \<Rightarrow> nat list => int"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
where
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
  "pflat_len v p \<equiv> (if p \<in> Pos v then intlen (flat (at v p)) else -1)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
lemma pflat_len_simps:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
  shows "pflat_len (Seq v1 v2) (0#p) = pflat_len v1 p"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
  and   "pflat_len (Seq v1 v2) (Suc 0#p) = pflat_len v2 p"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
  and   "pflat_len (Left v) (0#p) = pflat_len v p"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
  and   "pflat_len (Left v) (Suc 0#p) = -1"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
  and   "pflat_len (Right v) (Suc 0#p) = pflat_len v p"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
  and   "pflat_len (Right v) (0#p) = -1"
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    77
  and   "pflat_len (Stars (v#vs)) (Suc n#p) = pflat_len (Stars vs) (n#p)"
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    78
  and   "pflat_len (Stars (v#vs)) (0#p) = pflat_len v p"
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
  and   "pflat_len v [] = intlen (flat v)"
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    80
by (auto simp add: pflat_len_def Pos_empty)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
lemma pflat_len_Stars_simps:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
  assumes "n < length vs"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  shows "pflat_len (Stars vs) (n#p) = pflat_len (vs!n) p"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
using assms 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
apply(induct vs arbitrary: n p)
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
    87
apply(auto simp add: less_Suc_eq_0_disj pflat_len_simps)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    90
lemma outside_lemma:
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    91
  assumes "p \<notin> Pos v1 \<union> Pos v2"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    92
  shows "pflat_len v1 p = pflat_len v2 p"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    93
using assms by (auto simp add: pflat_len_def)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
    94
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
    95
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
    96
section {* Orderings *}
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
    97
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
    98
257
9deaff82e0c5 updated
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
    99
definition prefix_list:: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<sqsubseteq>pre _" [60,59] 60)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
where
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   101
  "ps1 \<sqsubseteq>pre ps2 \<equiv> \<exists>ps'. ps1 @ps' = ps2"
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
257
9deaff82e0c5 updated
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   103
definition sprefix_list:: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" ("_ \<sqsubset>spre _" [60,59] 60)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
where
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   105
  "ps1 \<sqsubset>spre ps2 \<equiv> ps1 \<sqsubseteq>pre ps2 \<and> ps1 \<noteq> ps2"
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
257
9deaff82e0c5 updated
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   107
inductive lex_list :: "nat list \<Rightarrow> nat list \<Rightarrow> bool" ("_ \<sqsubset>lex _" [60,59] 60)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
where
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   109
  "[] \<sqsubset>lex (p#ps)"
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
| "ps1 \<sqsubset>lex ps2 \<Longrightarrow> (p#ps1) \<sqsubset>lex (p#ps2)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
| "p1 < p2 \<Longrightarrow> (p1#ps1) \<sqsubset>lex (p2#ps2)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
lemma lex_irrfl:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
  fixes ps1 ps2 :: "nat list"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
  assumes "ps1 \<sqsubset>lex ps2"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
  shows "ps1 \<noteq> ps2"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
using assms
257
9deaff82e0c5 updated
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   118
by(induct rule: lex_list.induct)(auto)
9deaff82e0c5 updated
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   119
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   120
lemma lex_simps [simp]:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
  fixes xs ys :: "nat list"
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   122
  shows "[] \<sqsubset>lex ys \<longleftrightarrow> ys \<noteq> []"
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   123
  and   "xs \<sqsubset>lex [] \<longleftrightarrow> False"
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  and   "(x # xs) \<sqsubset>lex (y # ys) \<longleftrightarrow> (x < y \<or> (\<not> y < x \<and> xs \<sqsubset>lex ys))"
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   125
apply(auto elim: lex_list.cases intro: lex_list.intros)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   126
apply(auto simp add: neq_Nil_conv not_less_iff_gr_or_eq intro: lex_list.intros)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   127
done
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
lemma lex_trans:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
  fixes ps1 ps2 ps3 :: "nat list"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  assumes "ps1 \<sqsubset>lex ps2" "ps2 \<sqsubset>lex ps3"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
  shows "ps1 \<sqsubset>lex ps3"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
using assms
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   134
by (induct arbitrary: ps3 rule: lex_list.induct)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   135
   (auto elim: lex_list.cases)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   136
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
257
9deaff82e0c5 updated
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   138
lemma lex_trichotomous:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
  fixes p q :: "nat list"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
  shows "p = q \<or> p \<sqsubset>lex q \<or> q \<sqsubset>lex p"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
apply(induct p arbitrary: q)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
apply(auto)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
apply(case_tac q)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
apply(auto)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   147
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   148
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   149
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   150
section {* Ordering of values according to Okui & Suzuki *}
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   151
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   153
definition PosOrd:: "val \<Rightarrow> nat list \<Rightarrow> val \<Rightarrow> bool" ("_ \<sqsubset>val _ _" [60, 60, 59] 60)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   154
where
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   155
  "v1 \<sqsubset>val p v2 \<equiv> pflat_len v1 p > pflat_len v2 p \<and>
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   156
                   (\<forall>q \<in> Pos v1 \<union> Pos v2. q \<sqsubset>lex p \<longrightarrow> pflat_len v1 q = pflat_len v2 q)"
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   157
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   158
definition PosOrd2:: "val \<Rightarrow> nat list \<Rightarrow> val \<Rightarrow> bool" ("_ \<sqsubset>val2 _ _" [60, 60, 59] 60)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   159
where
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   160
  "v1 \<sqsubset>val2 p v2 \<equiv> p \<in> Pos v1 \<and> 
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   161
                   pflat_len v1 p > pflat_len v2 p \<and>
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   162
                   (\<forall>q \<in> Pos v1 \<union> Pos v2. q \<sqsubset>lex p \<longrightarrow> pflat_len v1 q = pflat_len v2 q)"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   163
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   164
lemma XXX:
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   165
  "v1 \<sqsubset>val p v2 \<longleftrightarrow> v1 \<sqsubset>val2 p v2"
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   166
apply(auto simp add: PosOrd_def PosOrd2_def)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   167
apply(auto simp add: pflat_len_def split: if_splits)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   168
by (smt intlen_bigger)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   169
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   170
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   171
(* some tests *)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   172
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   173
definition ValFlat_ord:: "val \<Rightarrow> nat list \<Rightarrow> val \<Rightarrow> bool" ("_ \<sqsubset>fval _ _")
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
where
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   175
  "v1 \<sqsubset>fval p v2 \<equiv> p \<in> Pos v1 \<and> 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   176
                    (p \<notin> Pos v2 \<or> flat (at v2 p) \<sqsubset>spre flat (at v1 p)) \<and>
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   177
                    (\<forall>q \<in> Pos v1 \<union> Pos v2. q \<sqsubset>lex p \<longrightarrow> (pflat_len v1 q = pflat_len v2 q))"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   178
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   179
lemma 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   180
  assumes "v1 \<sqsubset>fval p v2" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   181
  shows "v1 \<sqsubset>val p v2"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   182
using assms
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   183
unfolding ValFlat_ord_def PosOrd_def
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   184
apply(clarify)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   185
apply(simp add: pflat_len_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   186
apply(auto)[1]
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   187
apply(smt intlen_bigger)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   188
apply(simp add: sprefix_list_def prefix_list_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   189
apply(auto)[1]
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   190
apply(drule sym)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   191
apply(simp add: intlen_append)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   192
apply (metis intlen.simps(1) intlen_length length_greater_0_conv list.size(3))
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   193
apply(smt intlen_bigger)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   194
done
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   195
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   196
lemma 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   197
  assumes "v1 \<sqsubset>val p v2" "flat (at v2 p) \<sqsubset>spre flat (at v1 p)"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   198
  shows "v1 \<sqsubset>fval p v2"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   199
using assms
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   200
unfolding ValFlat_ord_def PosOrd_def
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   201
apply(clarify)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   202
oops
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   205
definition PosOrd_ex:: "val \<Rightarrow> val \<Rightarrow> bool" ("_ :\<sqsubset>val _" [60, 59] 60)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
where
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
  "v1 :\<sqsubset>val v2 \<equiv> (\<exists>p. v1 \<sqsubset>val p v2)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   209
definition PosOrd_ex1:: "val \<Rightarrow> val \<Rightarrow> bool" ("_ :\<sqsubseteq>val _" [60, 59] 60)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
where
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  "v1 :\<sqsubseteq>val v2 \<equiv> v1 :\<sqsubset>val v2 \<or> v1 = v2"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   213
lemma PosOrd_shorterE:
255
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   214
  assumes "v1 :\<sqsubset>val v2" 
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   215
  shows "length (flat v2) \<le> length (flat v1)"
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   216
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   217
apply(auto simp add: PosOrd_ex_def PosOrd_def)
255
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   218
apply(case_tac p)
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   219
apply(simp add: pflat_len_simps)
257
9deaff82e0c5 updated
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   220
apply(simp add: intlen_length)
255
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   221
apply(simp)
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   222
apply(drule_tac x="[]" in bspec)
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   223
apply(simp add: Pos_empty)
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   224
apply(simp add: pflat_len_simps)
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   225
by (metis intlen_length le_less less_irrefl linear)
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   226
222ed43892ea updated
Christian Urban <urbanc@in.tum.de>
parents: 254
diff changeset
   227
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   228
lemma PosOrd_shorterI:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
  assumes "length (flat v') < length (flat v)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
  shows "v :\<sqsubset>val v'" 
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   231
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   232
unfolding PosOrd_ex_def
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   233
by (metis intlen_length lex_simps(2) pflat_len_simps(9) PosOrd_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   235
lemma PosOrd_spreI:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  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:
diff changeset
   237
  shows "v :\<sqsubset>val v'" 
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   238
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   239
apply(rule_tac PosOrd_shorterI)
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   240
by (metis append_eq_conv_conj le_less_linear prefix_list_def sprefix_list_def take_all)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   242
lemma PosOrd_Left_Right:
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   243
  assumes "flat v1 = flat v2"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   244
  shows "Left v1 :\<sqsubset>val Right v2" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   245
unfolding PosOrd_ex_def
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   246
apply(rule_tac x="[0]" in exI)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   247
using assms
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   248
apply(auto simp add: PosOrd_def pflat_len_simps)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   249
apply(smt intlen_bigger)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   250
done
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   251
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   252
lemma PosOrd_LeftI:
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   253
  assumes "v :\<sqsubset>val v'" "flat v = flat v'"
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   254
  shows "(Left v) :\<sqsubset>val (Left v')" 
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
using assms(1)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   256
unfolding PosOrd_ex_def
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   257
apply(auto)
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   258
apply(rule_tac x="0#p" in exI)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
using assms(2)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   260
apply(auto simp add: PosOrd_def pflat_len_simps)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   263
lemma PosOrd_RightI:
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   264
  assumes "v :\<sqsubset>val v'" "flat v = flat v'"
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   265
  shows "(Right v) :\<sqsubset>val (Right v')" 
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
using assms(1)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   267
unfolding PosOrd_ex_def
251
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   268
apply(auto)
925232418a15 polished
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   269
apply(rule_tac x="Suc 0#p" in exI)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
using assms(2)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   271
apply(auto simp add: PosOrd_def pflat_len_simps)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   274
lemma PosOrd_LeftE:
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   275
  assumes "(Left v1) :\<sqsubset>val (Left v2)"
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   276
  shows "v1 :\<sqsubset>val v2"
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   277
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   278
apply(simp add: PosOrd_ex_def)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   279
apply(erule exE)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   280
apply(case_tac p)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   281
apply(simp add: PosOrd_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
apply(auto simp add: pflat_len_simps)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   283
apply(rule_tac x="[]" in exI)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   284
apply(simp add: Pos_empty pflat_len_simps)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   285
apply(auto simp add: pflat_len_simps PosOrd_def)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   286
apply(case_tac a)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   287
apply(auto simp add: pflat_len_simps)[1]
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   288
apply(rule_tac x="list" in exI)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   289
apply(auto)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   290
apply(drule_tac x="0#q" in bspec)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   291
apply(simp)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   292
apply(simp add: pflat_len_simps)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   293
apply(drule_tac x="0#q" in bspec)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   294
apply(simp)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   295
apply(simp add: pflat_len_simps)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   296
apply(simp add: pflat_len_def)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   297
done
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   299
lemma PosOrd_RightE:
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   300
  assumes "(Right v1) :\<sqsubset>val (Right v2)"
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   301
  shows "v1 :\<sqsubset>val v2"
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   302
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   303
apply(simp add: PosOrd_ex_def)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   304
apply(erule exE)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   305
apply(case_tac p)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   306
apply(simp add: PosOrd_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
apply(auto simp add: pflat_len_simps)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   308
apply(rule_tac x="[]" in exI)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   309
apply(simp add: Pos_empty pflat_len_simps)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   310
apply(case_tac a)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   311
apply(simp add: pflat_len_def PosOrd_def)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   312
apply(case_tac nat)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   313
prefer 2
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   314
apply(simp add: pflat_len_def PosOrd_def)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   315
apply(auto simp add: pflat_len_simps PosOrd_def)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   316
apply(rule_tac x="list" in exI)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   317
apply(auto)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   318
apply(drule_tac x="Suc 0#q" in bspec)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   319
apply(simp)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   320
apply(simp add: pflat_len_simps)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   321
apply(drule_tac x="Suc 0#q" in bspec)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   322
apply(simp)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   323
apply(simp add: pflat_len_simps)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   324
done
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   325
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   327
lemma PosOrd_SeqI1:
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   328
  assumes "v1 :\<sqsubset>val v1'" "flat (Seq v1 v2) = flat (Seq v1' v2')"
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   329
  shows "(Seq v1 v2) :\<sqsubset>val (Seq v1' v2')" 
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
using assms(1)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   331
apply(subst (asm) PosOrd_ex_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   332
apply(subst (asm) PosOrd_def)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   333
apply(clarify)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   334
apply(subst PosOrd_ex_def)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   335
apply(rule_tac x="0#p" in exI)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   336
apply(subst PosOrd_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
apply(rule conjI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
apply(simp add: pflat_len_simps)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
apply(rule ballI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
apply(rule impI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
apply(simp only: Pos.simps)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
apply(simp add: pflat_len_simps)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
using assms(2)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
apply(auto simp add: pflat_len_simps)[2]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   349
lemma PosOrd_SeqI2:
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   350
  assumes "v2 :\<sqsubset>val v2'" "flat v2 = flat v2'"
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   351
  shows "(Seq v v2) :\<sqsubset>val (Seq v v2')" 
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
using assms(1)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   353
apply(subst (asm) PosOrd_ex_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   354
apply(subst (asm) PosOrd_def)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   355
apply(clarify)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   356
apply(subst PosOrd_ex_def)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   357
apply(rule_tac x="Suc 0#p" in exI)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   358
apply(subst PosOrd_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
apply(rule conjI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
apply(simp add: pflat_len_simps)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
apply(rule ballI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
apply(rule impI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
apply(simp only: Pos.simps)
252
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   364
apply(auto)[1]
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   365
apply(simp add: pflat_len_simps)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   366
using assms(2)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   367
apply(simp)
022b80503766 polished
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   368
apply(auto simp add: pflat_len_simps)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   371
lemma PosOrd_SeqE:
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   372
  assumes "(Seq v1 v2) :\<sqsubset>val (Seq v1' v2')" 
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   373
  shows "v1 :\<sqsubset>val v1' \<or> v2 :\<sqsubset>val v2'"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   374
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   375
apply(simp add: PosOrd_ex_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   376
apply(erule exE)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   377
apply(case_tac p)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   378
apply(simp add: PosOrd_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   379
apply(auto simp add: pflat_len_simps intlen_append)[1]
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   380
apply(rule_tac x="[]" in exI)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   381
apply(drule_tac x="[]" in spec)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   382
apply(simp add: Pos_empty pflat_len_simps)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   383
apply(case_tac a)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   384
apply(rule disjI1)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   385
apply(simp add: PosOrd_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   386
apply(auto simp add: pflat_len_simps intlen_append)[1]
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   387
apply(rule_tac x="list" in exI)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   388
apply(simp)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   389
apply(rule ballI)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   390
apply(rule impI)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   391
apply(drule_tac x="0#q" in bspec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   392
apply(simp)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   393
apply(simp add: pflat_len_simps)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   394
apply(case_tac nat)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   395
apply(rule disjI2)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   396
apply(simp add: PosOrd_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   397
apply(auto simp add: pflat_len_simps intlen_append)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   398
apply(rule_tac x="list" in exI)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   399
apply(simp add: Pos_empty)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   400
apply(rule ballI)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   401
apply(rule impI)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   402
apply(auto)[1]
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   403
apply(drule_tac x="Suc 0#q" in bspec)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   404
apply(simp)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   405
apply(simp add: pflat_len_simps)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   406
apply(drule_tac x="Suc 0#q" in bspec)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   407
apply(simp)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   408
apply(simp add: pflat_len_simps)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   409
apply(simp add: PosOrd_def pflat_len_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   410
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   411
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   412
lemma PosOrd_StarsI:
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   413
  assumes "v1 :\<sqsubset>val v2" "flat (Stars (v1#vs1)) = flat (Stars (v2#vs2))"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   414
  shows "(Stars (v1#vs1)) :\<sqsubset>val (Stars (v2#vs2))" 
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   415
using assms(1)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   416
apply(subst (asm) PosOrd_ex_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   417
apply(subst (asm) PosOrd_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   418
apply(clarify)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   419
apply(subst PosOrd_ex_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   420
apply(subst PosOrd_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   421
apply(rule_tac x="0#p" in exI)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   422
apply(simp add: pflat_len_Stars_simps pflat_len_simps)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   423
using assms(2)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   424
apply(simp add: pflat_len_simps intlen_append)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   425
apply(auto simp add: pflat_len_Stars_simps pflat_len_simps)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   426
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   427
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   428
lemma PosOrd_StarsI2:
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   429
  assumes "(Stars vs1) :\<sqsubset>val (Stars vs2)" "flat (Stars vs1) = flat (Stars vs2)"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   430
  shows "(Stars (v#vs1)) :\<sqsubset>val (Stars (v#vs2))" 
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   431
using assms(1)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   432
apply(subst (asm) PosOrd_ex_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   433
apply(subst (asm) PosOrd_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   434
apply(clarify)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   435
apply(subst PosOrd_ex_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   436
apply(subst PosOrd_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   437
apply(case_tac p)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   438
apply(simp add: pflat_len_simps)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   439
apply(rule_tac x="[]" in exI)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   440
apply(simp add: pflat_len_Stars_simps pflat_len_simps intlen_append)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   441
apply(rule_tac x="Suc a#list" in exI)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   442
apply(simp add: pflat_len_Stars_simps pflat_len_simps)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   443
using assms(2)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   444
apply(simp add: pflat_len_simps intlen_append)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   445
apply(auto simp add: pflat_len_Stars_simps pflat_len_simps)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   446
done
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   447
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   448
lemma PosOrd_Stars_appendI:
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   449
  assumes "Stars vs1 :\<sqsubset>val Stars vs2" "flat (Stars vs1) = flat (Stars vs2)"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   450
  shows "Stars (vs @ vs1) :\<sqsubset>val Stars (vs @ vs2)"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   451
using assms
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   452
apply(induct vs)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   453
apply(simp)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   454
apply(simp add: PosOrd_StarsI2)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   455
done
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   456
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   457
lemma PosOrd_StarsE2:
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   458
  assumes "Stars (v # vs1) :\<sqsubset>val Stars (v # vs2)"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   459
  shows "Stars vs1 :\<sqsubset>val Stars vs2"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   460
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   461
apply(subst (asm) PosOrd_ex_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   462
apply(erule exE)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   463
apply(case_tac p)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   464
apply(simp)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   465
apply(simp add: PosOrd_def pflat_len_simps intlen_append)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   466
apply(subst PosOrd_ex_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   467
apply(rule_tac x="[]" in exI)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   468
apply(simp add: PosOrd_def pflat_len_simps Pos_empty)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   469
apply(simp)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   470
apply(case_tac a)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   471
apply(clarify)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   472
apply(auto simp add: pflat_len_simps PosOrd_def pflat_len_def split: if_splits)[1]
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   473
apply(clarify)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   474
apply(simp add: PosOrd_ex_def)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   475
apply(rule_tac x="nat#list" in exI)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   476
apply(auto simp add: PosOrd_def pflat_len_simps intlen_append)[1]
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   477
apply(case_tac q)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   478
apply(simp add: PosOrd_def pflat_len_simps intlen_append)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   479
apply(clarify)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   480
apply(drule_tac x="Suc a # lista" in bspec)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   481
apply(simp)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   482
apply(auto simp add: PosOrd_def pflat_len_simps intlen_append)[1]
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   483
apply(case_tac q)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   484
apply(simp add: PosOrd_def pflat_len_simps intlen_append)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   485
apply(clarify)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   486
apply(drule_tac x="Suc a # lista" in bspec)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   487
apply(simp)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   488
apply(auto simp add: PosOrd_def pflat_len_simps intlen_append)[1]
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   489
done
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   490
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   491
lemma PosOrd_Stars_appendE:
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   492
  assumes "Stars (vs @ vs1) :\<sqsubset>val Stars (vs @ vs2)"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   493
  shows "Stars vs1 :\<sqsubset>val Stars vs2"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   494
using assms
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   495
apply(induct vs)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   496
apply(simp)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   497
apply(simp add: PosOrd_StarsE2)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   498
done
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   499
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   500
lemma PosOrd_Stars_append_eq:
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   501
  assumes "flat (Stars vs1) = flat (Stars vs2)"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   502
  shows "Stars (vs @ vs1) :\<sqsubset>val Stars (vs @ vs2) \<longleftrightarrow> Stars vs1 :\<sqsubset>val Stars vs2"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   503
using assms
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   504
apply(rule_tac iffI)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   505
apply(erule PosOrd_Stars_appendE)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   506
apply(rule PosOrd_Stars_appendI)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   507
apply(auto)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   508
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   509
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   510
lemma PosOrd_trans:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   511
  assumes "v1 :\<sqsubset>val v2" "v2 :\<sqsubset>val v3"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   512
  shows "v1 :\<sqsubset>val v3"
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   513
proof -
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   514
  from assms obtain p p'
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   515
    where as: "v1 \<sqsubset>val p v2" "v2 \<sqsubset>val p' v3" unfolding PosOrd_ex_def by blast
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   516
  have "p = p' \<or> p \<sqsubset>lex p' \<or> p' \<sqsubset>lex p"
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   517
    by (rule lex_trichotomous)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   518
  moreover
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   519
    { assume "p = p'"
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   520
      with as have "v1 \<sqsubset>val p v3" unfolding PosOrd_def
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   521
      by (smt outside_lemma)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   522
      then have " v1 :\<sqsubset>val v3" unfolding PosOrd_ex_def by blast
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   523
    }   
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   524
  moreover
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   525
    { assume "p \<sqsubset>lex p'"
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   526
      with as have "v1 \<sqsubset>val p v3" unfolding PosOrd_def
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   527
      by (metis lex_trans outside_lemma)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   528
      then have " v1 :\<sqsubset>val v3" unfolding PosOrd_ex_def by blast
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   529
    }
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   530
  moreover
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   531
    { assume "p' \<sqsubset>lex p"
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   532
      with as have "v1 \<sqsubset>val p' v3" unfolding PosOrd_def
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   533
      by (smt Un_iff intlen_bigger lex_trans pflat_len_def)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   534
      then have "v1 :\<sqsubset>val v3" unfolding PosOrd_ex_def by blast
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   535
    }
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   536
  ultimately show "v1 :\<sqsubset>val v3" by blast
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   537
qed
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   538
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   539
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   540
lemma PosOrd_irrefl:
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   541
  assumes "v :\<sqsubset>val v"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   542
  shows "False"
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   543
using assms unfolding PosOrd_ex_def PosOrd_def
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   544
by auto
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   545
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   546
lemma PosOrd_almost_trichotomous:
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   547
  shows "v1 :\<sqsubset>val v2 \<or> v2 :\<sqsubset>val v1 \<or> (intlen (flat v1) = intlen (flat v2))"
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   548
apply(auto simp add: PosOrd_ex_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   549
apply(auto simp add: PosOrd_def)
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   550
apply(rule_tac x="[]" in exI)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   551
apply(auto simp add: Pos_empty pflat_len_simps)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   552
apply(drule_tac x="[]" in spec)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   553
apply(auto simp add: Pos_empty pflat_len_simps)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   554
done
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   555
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   556
lemma WW1:
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   557
  assumes "v1 :\<sqsubset>val v2" "v2 :\<sqsubset>val v1"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   558
  shows "False"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   559
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   560
apply(auto simp add: PosOrd_ex_def PosOrd_def)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   561
using assms PosOrd_irrefl PosOrd_trans by blast
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   562
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   563
lemma WW2:
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   564
  assumes "\<not>(v1 :\<sqsubset>val v2)" 
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   565
  shows "v1 = v2 \<or> v2 :\<sqsubset>val v1"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   566
using assms
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   567
oops
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   568
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   569
lemma PosOrd_SeqE2:
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   570
  assumes "(Seq v1 v2) :\<sqsubset>val (Seq v1' v2')" 
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   571
  shows "v1 :\<sqsubset>val v1' \<or> (v1 = v1' \<and> v2 :\<sqsubset>val v2')"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   572
using assms 
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   573
apply(frule_tac PosOrd_SeqE)
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   574
apply(erule disjE)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   575
apply(simp)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   576
apply(auto)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   577
apply(case_tac "v1 :\<sqsubset>val v1'")
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   578
apply(simp)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   579
apply(auto simp add: PosOrd_ex_def)
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   580
apply(case_tac "v1 = v1'")
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   581
apply(simp)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   582
oops
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   583
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   584
section {* CPT and CPTpre *}
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   585
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   586
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   587
inductive 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   588
  CPrf :: "val \<Rightarrow> rexp \<Rightarrow> bool" ("\<Turnstile> _ : _" [100, 100] 100)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   589
where
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   590
 "\<lbrakk>\<Turnstile> v1 : r1; \<Turnstile> v2 : r2\<rbrakk> \<Longrightarrow> \<Turnstile>  Seq v1 v2 : SEQ r1 r2"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   591
| "\<Turnstile> v1 : r1 \<Longrightarrow> \<Turnstile> Left v1 : ALT r1 r2"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   592
| "\<Turnstile> v2 : r2 \<Longrightarrow> \<Turnstile> Right v2 : ALT r1 r2"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   593
| "\<Turnstile> Void : ONE"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   594
| "\<Turnstile> Char c : CHAR c"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   595
| "\<Turnstile> Stars [] : STAR r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   596
| "\<lbrakk>\<Turnstile> v : r; flat v \<noteq> []; \<Turnstile> Stars vs : STAR r\<rbrakk> \<Longrightarrow> \<Turnstile> Stars (v # vs) : STAR r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   597
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   598
lemma Prf_CPrf:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   599
  assumes "\<Turnstile> v : r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   600
  shows "\<turnstile> v : r"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   601
using assms
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   602
by (induct) (auto intro: Prf.intros)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   603
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   604
lemma pflat_len_equal_iff:
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   605
  assumes "\<Turnstile> v1 : r" "\<Turnstile> v2 : r"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   606
  and "\<forall>p. pflat_len v1 p = pflat_len v2 p"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   607
  shows "v1 = v2"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   608
using assms
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   609
apply(induct v1 r arbitrary: v2 rule: CPrf.induct)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   610
apply(rotate_tac 4)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   611
apply(erule CPrf.cases)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   612
apply(simp_all)[7]
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   613
apply (metis pflat_len_simps(1) pflat_len_simps(2))
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   614
apply(rotate_tac 2)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   615
apply(erule CPrf.cases)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   616
apply(simp_all)[7]
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   617
apply (metis pflat_len_simps(3))
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   618
apply (metis intlen.simps(1) intlen_length length_greater_0_conv list.size(3) neg_0_le_iff_le not_less not_less_iff_gr_or_eq not_one_le_zero pflat_len_simps(3) pflat_len_simps(6) pflat_len_simps(9))
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   619
apply(rotate_tac 2)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   620
apply(erule CPrf.cases)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   621
apply(simp_all)[7]
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   622
apply (metis intlen.simps(1) intlen_length length_greater_0_conv list.size(3) neg_0_le_iff_le not_less not_less_iff_gr_or_eq not_one_le_zero pflat_len_simps(3) pflat_len_simps(6) pflat_len_simps(9))
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   623
apply (metis pflat_len_simps(5))
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   624
apply(erule CPrf.cases)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   625
apply(simp_all)[7]
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   626
apply(erule CPrf.cases)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   627
apply(simp_all)[7]
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   628
apply(erule CPrf.cases)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   629
apply(simp_all)[7]
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   630
apply (metis append_Cons flat.simps(6) flat.simps(7) intlen_length length_greater_0_conv neq_Nil_conv not_less_iff_gr_or_eq pflat_len_simps(9))
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   631
apply(rotate_tac 5)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   632
apply(erule CPrf.cases)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   633
apply(simp_all)[7]
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   634
apply (metis append_Cons flat.simps(6) flat.simps(7) intlen_length length_greater_0_conv list.distinct(1) list.exhaust not_less_iff_gr_or_eq pflat_len_simps(9))
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   635
apply(auto)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   636
apply (metis pflat_len_simps(8))
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   637
apply(subgoal_tac "v = va")
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   638
prefer 2
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   639
apply (metis pflat_len_simps(8))
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   640
apply(simp)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   641
apply(rotate_tac 3)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   642
apply(drule_tac x="Stars vsa" in meta_spec)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   643
apply(simp)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   644
apply(drule_tac meta_mp)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   645
apply(rule allI)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   646
apply(case_tac p)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   647
apply(simp add: pflat_len_simps)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   648
apply(drule_tac x="[]" in spec)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   649
apply(simp add: pflat_len_simps intlen_append)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   650
apply(drule_tac x="Suc a#list" in spec)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   651
apply(simp add: pflat_len_simps intlen_append)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   652
apply(simp)
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   653
done
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   654
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   655
lemma PosOrd_trichotomous_stronger:
256
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   656
  assumes "\<Turnstile> v1 : r" "\<Turnstile> v2 : r"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   657
  shows "v1 :\<sqsubset>val v2 \<or> v2 :\<sqsubset>val v1 \<or> (v1 = v2)"
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   658
oops
146b4817aebd updated
Christian Urban <urbanc@in.tum.de>
parents: 255
diff changeset
   659
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   660
lemma CPrf_stars:
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   661
  assumes "\<Turnstile> Stars vs : STAR r"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   662
  shows "\<forall>v \<in> set vs. flat v \<noteq> [] \<and> \<Turnstile> v : r"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   663
using assms
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   664
apply(induct vs)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   665
apply(auto)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   666
apply(erule CPrf.cases)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   667
apply(simp_all)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   668
apply(erule CPrf.cases)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   669
apply(simp_all)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   670
apply(erule CPrf.cases)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   671
apply(simp_all)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   672
apply(erule CPrf.cases)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   673
apply(simp_all)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   674
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   675
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   676
lemma CPrf_Stars_appendE:
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   677
  assumes "\<Turnstile> Stars (vs1 @ vs2) : STAR r"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   678
  shows "\<Turnstile> Stars vs1 : STAR r \<and> \<Turnstile> Stars vs2 : STAR r" 
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   679
using assms
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   680
apply(induct vs1 arbitrary: vs2)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   681
apply(auto intro: CPrf.intros)[1]
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   682
apply(erule CPrf.cases)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   683
apply(simp_all)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   684
apply(auto intro: CPrf.intros)
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   685
done
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   686
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   687
definition PT :: "rexp \<Rightarrow> string \<Rightarrow> val set"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   688
where "PT r s \<equiv> {v. flat v = s \<and> \<turnstile> v : r}"
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   689
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   690
definition
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   691
  "CPT r s = {v. flat v = s \<and> \<Turnstile> v : r}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   692
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   693
definition
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   694
  "CPTpre r s = {v. \<exists>s'. flat v @ s' = s \<and> \<Turnstile> v : r}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   695
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   696
lemma CPT_CPTpre_subset:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   697
  shows "CPT r s \<subseteq> CPTpre r s"
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
   698
by(auto simp add: CPT_def CPTpre_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   699
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   700
lemma CPTpre_subsets:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   701
  "CPTpre ZERO s = {}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   702
  "CPTpre ONE s \<subseteq> {Void}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   703
  "CPTpre (CHAR c) s \<subseteq> {Char c}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   704
  "CPTpre (ALT r1 r2) s \<subseteq> Left ` CPTpre r1 s \<union> Right ` CPTpre r2 s"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   705
  "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)}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   706
  "CPTpre (STAR r) s \<subseteq> {Stars []} \<union>
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   707
    {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)}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   708
  "CPTpre (STAR r) [] = {Stars []}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   709
apply(auto simp add: CPTpre_def)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   710
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   711
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   712
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   713
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   714
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   715
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   716
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   717
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   718
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   719
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   720
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   721
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   722
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   723
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   724
apply(rule CPrf.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   725
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   726
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   727
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   728
lemma CPTpre_simps:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   729
  shows "CPTpre ONE s = {Void}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   730
  and   "CPTpre (CHAR c) (d#s) = (if c = d then {Char c} else {})"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   731
  and   "CPTpre (ALT r1 r2) s = Left ` CPTpre r1 s \<union> Right ` CPTpre r2 s"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   732
  and   "CPTpre (SEQ r1 r2) s = 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   733
          {Seq v1 v2 | v1 v2. v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   734
apply -
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   735
apply(rule subset_antisym)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   736
apply(rule CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   737
apply(auto simp add: CPTpre_def intro: "CPrf.intros")[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   738
apply(case_tac "c = d")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   739
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   740
apply(rule subset_antisym)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   741
apply(rule CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   742
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   743
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   744
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   745
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   746
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   747
apply(rule subset_antisym)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   748
apply(rule CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   749
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   750
apply(rule subset_antisym)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   751
apply(rule CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   752
apply(auto simp add: CPTpre_def intro: CPrf.intros)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   753
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   754
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   755
lemma CPT_simps:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   756
  shows "CPT ONE s = (if s = [] then {Void} else {})"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   757
  and   "CPT (CHAR c) [d] = (if c = d then {Char c} else {})"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   758
  and   "CPT (ALT r1 r2) s = Left ` CPT r1 s \<union> Right ` CPT r2 s"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   759
  and   "CPT (SEQ r1 r2) s = 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   760
          {Seq v1 v2 | v1 v2 s1 s2. s1 @ s2 = s \<and> v1 \<in> CPT r1 s1 \<and> v2 \<in> CPT r2 s2}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   761
apply -
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   762
apply(rule subset_antisym)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   763
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:
diff changeset
   764
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   765
apply(simp_all)[7]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   766
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   767
apply(simp_all)[7]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   768
apply(auto simp add: CPT_def intro: CPrf.intros)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   769
apply(auto simp add: CPT_def intro: CPrf.intros)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   770
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   771
apply(simp_all)[7]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   772
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   773
apply(simp_all)[7]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   774
apply(auto simp add: CPT_def image_def intro: CPrf.intros)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   775
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   776
apply(simp_all)[7]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   777
apply(clarify)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   778
apply blast
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   779
apply(auto simp add: CPT_def image_def intro: CPrf.intros)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   780
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   781
apply(simp_all)[7]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   782
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   783
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   784
lemma test: 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   785
  assumes "finite A"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   786
  shows "finite {vs. Stars vs \<in> A}"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   787
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   788
apply(induct A)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   789
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   790
apply(auto)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   791
apply(case_tac x)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   792
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   793
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   794
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   795
lemma CPTpre_STAR_finite:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   796
  assumes "\<And>s. finite (CPTpre r s)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   797
  shows "finite (CPTpre (STAR r) s)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   798
apply(induct s rule: length_induct)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   799
apply(case_tac xs)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   800
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   801
apply(simp add: CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   802
apply(rule finite_subset)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   803
apply(rule CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   804
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   805
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)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   806
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   807
apply(rule finite_imageI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   808
apply(simp add: Collect_case_prod_Sigma)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   809
apply(rule finite_SigmaI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   810
apply(rule assms)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   811
apply(case_tac "flat v = []")
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   812
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   813
apply(drule_tac x="drop (length (flat v)) (a # list)" in spec)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   814
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   815
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   816
apply(rule test)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   817
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   818
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   819
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   820
lemma CPTpre_finite:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   821
  shows "finite (CPTpre r s)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   822
apply(induct r arbitrary: s)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   823
apply(simp add: CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   824
apply(rule finite_subset)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   825
apply(rule CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   826
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   827
apply(rule finite_subset)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   828
apply(rule CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   829
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   830
apply(rule finite_subset)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   831
apply(rule CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   832
apply(rule_tac B="(\<lambda>(v1, v2). Seq v1 v2) ` {(v1, v2).  v1 \<in> CPTpre r1 s \<and> v2 \<in> CPTpre r2 (drop (length (flat v1)) s)}" in finite_subset)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   833
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   834
apply(rule finite_imageI)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   835
apply(simp add: Collect_case_prod_Sigma)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   836
apply(rule finite_subset)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   837
apply(rule CPTpre_subsets)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   838
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   839
by (simp add: CPTpre_STAR_finite)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   840
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   841
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   842
lemma CPT_finite:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   843
  shows "finite (CPT r s)"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   844
apply(rule finite_subset)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   845
apply(rule CPT_CPTpre_subset)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   846
apply(rule CPTpre_finite)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   847
done
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   848
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   849
lemma Posix_CPT:
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   850
  assumes "s \<in> r \<rightarrow> v"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   851
  shows "v \<in> CPT r s"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   852
using assms
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   853
by (induct rule: Posix.induct)
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   854
   (auto simp add: CPT_def intro: CPrf.intros)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   855
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   856
section {* The Posix Value is smaller than any other Value *}
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   857
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   858
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   859
lemma Posix_PosOrd:
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   860
  assumes "s \<in> r \<rightarrow> v1" "v2 \<in> CPT r s" 
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   861
  shows "v1 :\<sqsubseteq>val v2"
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   862
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   863
proof (induct arbitrary: v2 rule: Posix.induct)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   864
  case (Posix_ONE v)
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   865
  have "v \<in> CPT ONE []" by fact
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   866
  then have "v = Void"
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   867
    by (simp add: CPT_simps)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   868
  then show "Void :\<sqsubseteq>val v"
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   869
    by (simp add: PosOrd_ex1_def)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   870
next
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   871
  case (Posix_CHAR c v)
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   872
  have "v \<in> CPT (CHAR c) [c]" by fact
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   873
  then have "v = Char c"
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   874
    by (simp add: CPT_simps)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   875
  then show "Char c :\<sqsubseteq>val v"
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   876
    by (simp add: PosOrd_ex1_def)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   877
next
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   878
  case (Posix_ALT1 s r1 v r2 v2)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   879
  have as1: "s \<in> r1 \<rightarrow> v" by fact
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   880
  have IH: "\<And>v2. v2 \<in> CPT r1 s \<Longrightarrow> v :\<sqsubseteq>val v2" by fact
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   881
  have "v2 \<in> CPT (ALT r1 r2) s" by fact
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   882
  then have "\<Turnstile> v2 : ALT r1 r2" "flat v2 = s"
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   883
    by(auto simp add: CPT_def prefix_list_def)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   884
  then consider
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   885
    (Left) v3 where "v2 = Left v3" "\<Turnstile> v3 : r1" "flat v3 = s" 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   886
  | (Right) v3 where "v2 = Right v3" "\<Turnstile> v3 : r2" "flat v3 = s"
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   887
  by (auto elim: CPrf.cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   888
  then show "Left v :\<sqsubseteq>val v2"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   889
  proof(cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   890
     case (Left v3)
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   891
     have "v3 \<in> CPT r1 s" using Left(2,3) 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   892
       by (auto simp add: CPT_def prefix_list_def)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   893
     with IH have "v :\<sqsubseteq>val v3" by simp
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   894
     moreover
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   895
     have "flat v3 = flat v" using as1 Left(3)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   896
       by (simp add: Posix1(2)) 
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   897
     ultimately have "Left v :\<sqsubseteq>val Left v3"
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   898
       by (auto simp add: PosOrd_ex1_def PosOrd_LeftI)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   899
     then show "Left v :\<sqsubseteq>val v2" unfolding Left .
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   900
  next
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   901
     case (Right v3)
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   902
     have "flat v3 = flat v" using as1 Right(3)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   903
       by (simp add: Posix1(2)) 
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   904
     then have "Left v :\<sqsubseteq>val Right v3" using Right(3) as1 
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   905
       by (auto simp add: PosOrd_ex1_def PosOrd_Left_Right)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   906
     then show "Left v :\<sqsubseteq>val v2" unfolding Right .
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   907
  qed
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   908
next
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   909
  case (Posix_ALT2 s r2 v r1 v2)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   910
  have as1: "s \<in> r2 \<rightarrow> v" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   911
  have as2: "s \<notin> L r1" by fact
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   912
  have IH: "\<And>v2. v2 \<in> CPT r2 s \<Longrightarrow> v :\<sqsubseteq>val v2" by fact
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   913
  have "v2 \<in> CPT (ALT r1 r2) s" by fact
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   914
  then have "\<Turnstile> v2 : ALT r1 r2" "flat v2 = s"
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   915
    by(auto simp add: CPT_def prefix_list_def)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   916
  then consider
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   917
    (Left) v3 where "v2 = Left v3" "\<Turnstile> v3 : r1" "flat v3 = s" 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   918
  | (Right) v3 where "v2 = Right v3" "\<Turnstile> v3 : r2" "flat v3 = s"
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   919
  by (auto elim: CPrf.cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   920
  then show "Right v :\<sqsubseteq>val v2"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   921
  proof (cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   922
    case (Right v3)
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   923
     have "v3 \<in> CPT r2 s" using Right(2,3) 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   924
       by (auto simp add: CPT_def prefix_list_def)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   925
     with IH have "v :\<sqsubseteq>val v3" by simp
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   926
     moreover
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   927
     have "flat v3 = flat v" using as1 Right(3)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   928
       by (simp add: Posix1(2)) 
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   929
     ultimately have "Right v :\<sqsubseteq>val Right v3" 
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   930
        by (auto simp add: PosOrd_ex1_def PosOrd_RightI)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   931
     then show "Right v :\<sqsubseteq>val v2" unfolding Right .
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   932
  next
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   933
     case (Left v3)
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   934
     have "v3 \<in> CPT r1 s" using Left(2,3) as2  
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   935
       by (auto simp add: CPT_def prefix_list_def)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   936
     then have "flat v3 = flat v \<and> \<Turnstile> v3 : r1" using as1 Left(3)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   937
       by (simp add: Posix1(2) CPT_def) 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   938
     then have "False" using as1 as2 Left
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   939
       by (auto simp add: Posix1(2) L_flat_Prf1 Prf_CPrf)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   940
     then show "Right v :\<sqsubseteq>val v2" by simp
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   941
  qed
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   942
next 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   943
  case (Posix_SEQ s1 r1 v1 s2 r2 v2 v3)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   944
  have "s1 \<in> r1 \<rightarrow> v1" "s2 \<in> r2 \<rightarrow> v2" by fact+
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   945
  then have as1: "s1 = flat v1" "s2 = flat v2" by (simp_all add: Posix1(2))
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   946
  have IH1: "\<And>v3. v3 \<in> CPT r1 s1 \<Longrightarrow> v1 :\<sqsubseteq>val v3" by fact
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   947
  have IH2: "\<And>v3. v3 \<in> CPT r2 s2 \<Longrightarrow> v2 :\<sqsubseteq>val v3" by fact
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   948
  have cond: "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" by fact
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   949
  have "v3 \<in> CPT (SEQ r1 r2) (s1 @ s2)" by fact
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   950
  then obtain v3a v3b where eqs:
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   951
    "v3 = Seq v3a v3b" "\<Turnstile> v3a : r1" "\<Turnstile> v3b : r2"
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   952
    "flat v3a @ flat v3b = s1 @ s2" 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   953
    by (force simp add: prefix_list_def CPT_def elim: CPrf.cases)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   954
  with cond have "flat v3a \<sqsubseteq>pre s1" unfolding prefix_list_def
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   955
    by (smt L_flat_Prf1 Prf_CPrf append_eq_append_conv2 append_self_conv)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   956
  then have "flat v3a \<sqsubset>spre s1 \<or> (flat v3a = s1 \<and> flat v3b = s2)" using eqs
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   957
    by (simp add: sprefix_list_def append_eq_conv_conj)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   958
  then have q2: "v1 :\<sqsubset>val v3a \<or> (flat v3a = s1 \<and> flat v3b = s2)" 
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   959
    using PosOrd_spreI as1(1) eqs by blast
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   960
  then have "v1 :\<sqsubset>val v3a \<or> (v3a \<in> CPT r1 s1 \<and> v3b \<in> CPT r2 s2)" using eqs(2,3)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   961
    by (auto simp add: CPT_def)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   962
  then have "v1 :\<sqsubset>val v3a \<or> (v1 :\<sqsubseteq>val v3a \<and> v2 :\<sqsubseteq>val v3b)" using IH1 IH2 by blast         
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   963
  then have "Seq v1 v2 :\<sqsubseteq>val Seq v3a v3b" using eqs q2 as1
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   964
    unfolding  PosOrd_ex1_def by (auto simp add: PosOrd_SeqI1 PosOrd_SeqI2) 
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   965
  then show "Seq v1 v2 :\<sqsubseteq>val v3" unfolding eqs by blast
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   966
next 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   967
  case (Posix_STAR1 s1 r v s2 vs v3) 
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   968
  have "s1 \<in> r \<rightarrow> v" "s2 \<in> STAR r \<rightarrow> Stars vs" by fact+
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   969
  then have as1: "s1 = flat v" "s2 = flat (Stars vs)" by (auto dest: Posix1(2))
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   970
  have IH1: "\<And>v3. v3 \<in> CPT r s1 \<Longrightarrow> v :\<sqsubseteq>val v3" by fact
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   971
  have IH2: "\<And>v3. v3 \<in> CPT (STAR r) s2 \<Longrightarrow> Stars vs :\<sqsubseteq>val v3" by fact
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   972
  have cond: "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = s2 \<and> s1 @ s\<^sub>3 \<in> L r \<and> s\<^sub>4 \<in> L (STAR r))" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   973
  have cond2: "flat v \<noteq> []" by fact
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   974
  have "v3 \<in> CPT (STAR r) (s1 @ s2)" by fact
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   975
  then consider
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   976
    (NonEmpty) v3a vs3 where
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   977
    "v3 = Stars (v3a # vs3)" "\<Turnstile> v3a : r" "\<Turnstile> Stars vs3 : STAR r"
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   978
    "flat (Stars (v3a # vs3)) = s1 @ s2"
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   979
  | (Empty) "v3 = Stars []"
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   980
  by (force simp add: CPT_def elim: CPrf.cases)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   981
  then show "Stars (v # vs) :\<sqsubseteq>val v3"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   982
    proof (cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
   983
      case (NonEmpty v3a vs3)
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   984
      have "flat (Stars (v3a # vs3)) = s1 @ s2" using NonEmpty(4) . 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   985
      with cond have "flat v3a \<sqsubseteq>pre s1" using NonEmpty(2,3)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   986
        unfolding prefix_list_def
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   987
        by (smt L_flat_Prf1 Prf_CPrf append_Nil2 append_eq_append_conv2 flat.simps(7)) 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   988
      then have "flat v3a \<sqsubset>spre s1 \<or> (flat v3a = s1 \<and> flat (Stars vs3) = s2)" using NonEmpty(4)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   989
        by (simp add: sprefix_list_def append_eq_conv_conj)
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   990
      then have q2: "v :\<sqsubset>val v3a \<or> (flat v3a = s1 \<and> flat (Stars vs3) = s2)" 
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   991
        using PosOrd_spreI as1(1) NonEmpty(4) by blast
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   992
      then have "v :\<sqsubset>val v3a \<or> (v3a \<in> CPT r s1 \<and> Stars vs3 \<in> CPT (STAR r) s2)" 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   993
        using NonEmpty(2,3) by (auto simp add: CPT_def)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   994
      then have "v :\<sqsubset>val v3a \<or> (v :\<sqsubseteq>val v3a \<and> Stars vs :\<sqsubseteq>val Stars vs3)" using IH1 IH2 by blast
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   995
      then have "v :\<sqsubset>val v3a \<or> (v = v3a \<and> Stars vs :\<sqsubseteq>val Stars vs3)" 
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   996
         unfolding PosOrd_ex1_def by auto     
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   997
      then have "Stars (v # vs) :\<sqsubseteq>val Stars (v3a # vs3)" using NonEmpty(4) q2 as1
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
   998
        unfolding  PosOrd_ex1_def
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   999
        by (metis PosOrd_StarsI PosOrd_StarsI2 flat.simps(7) val.inject(5))
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
  1000
      then show "Stars (v # vs) :\<sqsubseteq>val v3" unfolding NonEmpty by blast
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1001
    next 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1002
      case Empty
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1003
      have "v3 = Stars []" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1004
      then show "Stars (v # vs) :\<sqsubseteq>val v3"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1005
      unfolding PosOrd_ex1_def using cond2
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1006
      by (simp add: PosOrd_shorterI)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1007
    qed      
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1008
next 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1009
  case (Posix_STAR2 r v2)
262
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
  1010
  have "v2 \<in> CPT (STAR r) []" by fact
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
  1011
  then have "v2 = Stars []" 
45ad887fa6aa isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 261
diff changeset
  1012
    unfolding CPT_def by (auto elim: CPrf.cases) 
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1013
  then show "Stars [] :\<sqsubseteq>val v2"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1014
  by (simp add: PosOrd_ex1_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1015
qed
253
ca4e9eb8d576 polished
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
  1016
263
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1017
lemma Posix_PosOrd_stronger:
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1018
  assumes "s \<in> r \<rightarrow> v1" "v2 \<in> CPTpre r s" 
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1019
  shows "v1 :\<sqsubseteq>val v2"
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1020
proof -
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1021
  from assms(2) have "v2 \<in> CPT r s \<or> flat v2 \<sqsubset>spre s"
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1022
  unfolding CPTpre_def CPT_def sprefix_list_def prefix_list_def by auto
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1023
  moreover
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1024
    { assume "v2 \<in> CPT r s" 
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
  1025
      with assms(1) 
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
  1026
      have "v1 :\<sqsubseteq>val v2" by (rule Posix_PosOrd)
263
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1027
    }
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1028
  moreover
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1029
    { assume "flat v2 \<sqsubset>spre s"
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1030
      then have "flat v2 \<sqsubset>spre flat v1" using assms(1)
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1031
        using Posix1(2) by blast
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1032
      then have "v1 :\<sqsubseteq>val v2"
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1033
        by (simp add: PosOrd_ex1_def PosOrd_spreI) 
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1034
    }
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1035
  ultimately show "v1 :\<sqsubseteq>val v2" by blast
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1036
qed
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1037
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1038
lemma Posix_PosOrd_reverse:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1039
  assumes "s \<in> r \<rightarrow> v1" 
263
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1040
  shows "\<not>(\<exists>v2 \<in> CPTpre r s. v2 :\<sqsubset>val v1)"
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1041
using assms
263
00c9a95d492e isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1042
by (metis Posix_PosOrd_stronger less_irrefl PosOrd_def 
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1043
    PosOrd_ex1_def PosOrd_ex_def PosOrd_trans)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1044
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1045
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1046
lemma PosOrd_Posix_Stars:
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1047
  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:
diff changeset
  1048
  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:
diff changeset
  1049
  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:
diff changeset
  1050
using assms
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1051
apply(induct vs)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1052
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1053
apply(rule Posix.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1054
apply(simp (no_asm))
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1055
apply(rule Posix.intros)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1056
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1057
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:
diff changeset
  1058
defer
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1059
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1060
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1061
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:
diff changeset
  1062
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1063
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1064
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1065
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:
diff changeset
  1066
apply(erule Prf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1067
apply(simp_all)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
  1068
using CPrf_stars PosOrd_irrefl apply fastforce
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1069
apply(clarify)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1070
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:
diff changeset
  1071
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1072
apply(drule mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1073
apply (meson CPrf_stars Prf.intros(7) Prf_CPrf list.set_intros(1))
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1074
apply(subst (asm) (2) PosOrd_ex_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1075
apply(simp)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1076
apply (metis flat.simps(7) flat_Stars PosOrd_StarsI2 PosOrd_ex_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1077
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:
diff changeset
  1078
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:
diff changeset
  1079
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1080
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:
diff changeset
  1081
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:
diff changeset
  1082
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1083
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:
diff changeset
  1084
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:
diff changeset
  1085
apply(clarify)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1086
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:
diff changeset
  1087
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1088
apply(drule mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1089
using Prf.intros(7) apply blast
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1090
apply(subst (asm) (2) PosOrd_ex_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1091
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1092
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1093
apply(simp)
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
  1094
using Star_values_exists apply blast
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1095
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1096
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1097
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1098
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1099
apply(drule meta_mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1100
apply(auto)[1]
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1101
prefer 2
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1102
apply(simp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1103
apply(erule CPrf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1104
apply(simp_all)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1105
apply(clarify)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1106
apply(rotate_tac 3)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1107
apply(erule Prf.cases)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1108
apply(simp_all)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1109
apply (metis CPrf_stars intlen.cases less_irrefl list.set_intros(1) PosOrd_def PosOrd_ex_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1110
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:
diff changeset
  1111
apply(drule mp)
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1112
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:
diff changeset
  1113
apply(simp)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1114
apply(subst (asm) (2) PosOrd_ex_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1115
apply(simp)
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1116
apply (metis flat.simps(7) flat_Stars PosOrd_StarsI2 PosOrd_ex_def)
264
e2828c4a1e23 updated
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
  1117
by (metis PosOrd_StarsI PosOrd_ex_def PosOrd_spreI append_assoc append_self_conv flat.simps(7) flat_Stars prefix_list_def sprefix_list_def)
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1118
254
7c89d3f6923e polished
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
  1119
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1120
section {* The Smallest Value is indeed the Posix Value *}
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1121
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1122
lemma PosOrd_Posix:
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1123
  assumes "v1 \<in> CPT r s" "\<forall>v2 \<in> PT r s. \<not> v2 :\<sqsubset>val v1"
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1124
  shows "s \<in> r \<rightarrow> v1" 
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1125
using assms
261
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1126
proof(induct r arbitrary: s v1)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1127
  case (ZERO s v1)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1128
  have "v1 \<in> CPT ZERO s" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1129
  then show "s \<in> ZERO \<rightarrow> v1" unfolding CPT_def
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1130
    by (auto elim: CPrf.cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1131
next 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1132
  case (ONE s v1)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1133
  have "v1 \<in> CPT ONE s" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1134
  then show "s \<in> ONE \<rightarrow> v1" unfolding CPT_def
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1135
    by(auto elim!: CPrf.cases intro: Posix.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1136
next 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1137
  case (CHAR c s v1)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1138
  have "v1 \<in> CPT (CHAR c) s" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1139
  then show "s \<in> CHAR c \<rightarrow> v1" unfolding CPT_def
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1140
    by (auto elim!: CPrf.cases intro: Posix.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1141
next
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1142
  case (ALT r1 r2 s v1)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1143
  have IH1: "\<And>s v1. \<lbrakk>v1 \<in> CPT r1 s; \<forall>v2 \<in> PT r1 s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r1 \<rightarrow> v1" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1144
  have IH2: "\<And>s v1. \<lbrakk>v1 \<in> CPT r2 s; \<forall>v2 \<in> PT r2 s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r2 \<rightarrow> v1" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1145
  have as1: "\<forall>v2\<in>PT (ALT r1 r2) s. \<not> v2 :\<sqsubset>val v1" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1146
  have as2: "v1 \<in> CPT (ALT r1 r2) s" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1147
  then consider 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1148
     (Left) v1' where
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1149
        "v1 = Left v1'" "s = flat v1'"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1150
        "v1' \<in> CPT r1 s"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1151
  |  (Right) v1' where
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1152
        "v1 = Right v1'" "s = flat v1'"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1153
        "v1' \<in> CPT r2 s"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1154
  unfolding CPT_def by (auto elim: CPrf.cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1155
  then show "s \<in> ALT r1 r2 \<rightarrow> v1"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1156
   proof (cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1157
     case (Left v1')
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1158
     have "v1' \<in> CPT r1 s" using as2
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1159
       unfolding CPT_def Left by (auto elim: CPrf.cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1160
     moreover
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1161
     have "\<forall>v2 \<in> PT r1 s. \<not> v2 :\<sqsubset>val v1'" using as1
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1162
       unfolding PT_def Left using Prf.intros(2) PosOrd_LeftI by force  
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1163
     ultimately have "s \<in> r1 \<rightarrow> v1'" using IH1 by simp
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1164
     then have "s \<in> ALT r1 r2 \<rightarrow> Left v1'" by (rule Posix.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1165
     then show "s \<in> ALT r1 r2 \<rightarrow> v1" using Left by simp
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1166
   next
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1167
     case (Right v1')
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1168
     have "v1' \<in> CPT r2 s" using as2
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1169
       unfolding CPT_def Right by (auto elim: CPrf.cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1170
     moreover
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1171
     have "\<forall>v2 \<in> PT r2 s. \<not> v2 :\<sqsubset>val v1'" using as1
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1172
       unfolding PT_def Right using Prf.intros(3) PosOrd_RightI by force   
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1173
     ultimately have "s \<in> r2 \<rightarrow> v1'" using IH2 by simp
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1174
     moreover 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1175
       { assume "s \<in> L r1"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1176
         then obtain v' where "v' \<in>  PT r1 s"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1177
            unfolding PT_def using L_flat_Prf2 by blast
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1178
         then have "Left v' \<in>  PT (ALT r1 r2) s" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1179
            unfolding PT_def by (auto intro: Prf.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1180
         with as1 have "\<not> (Left v' :\<sqsubset>val Right v1') \<and> (flat v' = s)" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1181
            unfolding PT_def Right by (auto)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1182
         then have False using PosOrd_Left_Right Right by blast  
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1183
       }
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1184
     then have "s \<notin> L r1" by rule 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1185
     ultimately have "s \<in> ALT r1 r2 \<rightarrow> Right v1'" by (rule Posix.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1186
     then show "s \<in> ALT r1 r2 \<rightarrow> v1" using Right by simp
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1187
  qed
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1188
next 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1189
  case (SEQ r1 r2 s v1)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1190
  have IH1: "\<And>s v1. \<lbrakk>v1 \<in> CPT r1 s; \<forall>v2 \<in> PT r1 s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r1 \<rightarrow> v1" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1191
  have IH2: "\<And>s v1. \<lbrakk>v1 \<in> CPT r2 s; \<forall>v2 \<in> PT r2 s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r2 \<rightarrow> v1" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1192
  have as1: "\<forall>v2\<in>PT (SEQ r1 r2) s. \<not> v2 :\<sqsubset>val v1" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1193
  have as2: "v1 \<in> CPT (SEQ r1 r2) s" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1194
  then obtain 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1195
    v1a v1b where eqs:
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1196
        "v1 = Seq v1a v1b" "s = flat v1a @ flat v1b"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1197
        "v1a \<in> CPT r1 (flat v1a)" "v1b \<in> CPT r2 (flat v1b)" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1198
  unfolding CPT_def by(auto elim: CPrf.cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1199
  have "\<forall>v2 \<in> PT r1 (flat v1a). \<not> v2 :\<sqsubset>val v1a"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1200
    proof
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1201
      fix v2
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1202
      assume "v2 \<in> PT r1 (flat v1a)"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1203
      with eqs(2,4) have "Seq v2 v1b \<in> PT (SEQ r1 r2) s"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1204
         by (simp add: CPT_def PT_def Prf.intros(1) Prf_CPrf)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1205
      with as1 have "\<not> Seq v2 v1b :\<sqsubset>val Seq v1a v1b \<and> flat (Seq v2 v1b) = flat (Seq v1a v1b)" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1206
         using eqs by (simp add: PT_def) 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1207
      then show "\<not> v2 :\<sqsubset>val v1a"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1208
         using PosOrd_SeqI1 by blast
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1209
    qed     
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1210
  then have "flat v1a \<in> r1 \<rightarrow> v1a" using IH1 eqs by simp
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1211
  moreover
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1212
  have "\<forall>v2 \<in> PT r2 (flat v1b). \<not> v2 :\<sqsubset>val v1b"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1213
    proof 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1214
      fix v2
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1215
      assume "v2 \<in> PT r2 (flat v1b)"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1216
      with eqs(2,3,4) have "Seq v1a v2 \<in> PT (SEQ r1 r2) s"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1217
         by (simp add: CPT_def PT_def Prf.intros Prf_CPrf)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1218
      with as1 have "\<not> Seq v1a v2 :\<sqsubset>val Seq v1a v1b \<and> flat v2 = flat v1b" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1219
         using eqs by (simp add: PT_def) 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1220
      then show "\<not> v2 :\<sqsubset>val v1b"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1221
         using PosOrd_SeqI2 by auto
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1222
    qed     
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1223
  then have "flat v1b \<in> r2 \<rightarrow> v1b" using IH2 eqs by simp    
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1224
  moreover
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1225
  have "\<not> (\<exists>s\<^sub>3 s\<^sub>4. s\<^sub>3 \<noteq> [] \<and> s\<^sub>3 @ s\<^sub>4 = flat v1b \<and> flat v1a @ s\<^sub>3 \<in> L r1 \<and> s\<^sub>4 \<in> L r2)" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1226
  proof
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1227
     assume "\<exists>s3 s4. s3 \<noteq> [] \<and> s3 @ s4 = flat v1b \<and> flat v1a @ s3 \<in> L r1 \<and> s4 \<in> L r2"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1228
     then obtain s3 s4 where q1: "s3 \<noteq> [] \<and> s3 @ s4 = flat v1b \<and> flat v1a @ s3 \<in> L r1 \<and> s4 \<in> L r2" by blast
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1229
     then obtain vA vB where q2: "flat vA = flat v1a @ s3" "\<turnstile> vA : r1" "flat vB = s4" "\<turnstile> vB : r2"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1230
        using L_flat_Prf2 by blast
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1231
     then have "Seq vA vB \<in> PT (SEQ r1 r2) s" unfolding eqs using q1
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1232
       by (auto simp add: PT_def intro: Prf.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1233
     with as1 have "\<not> Seq vA vB :\<sqsubset>val Seq v1a v1b" unfolding eqs by auto
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1234
     then have "\<not> vA :\<sqsubset>val v1a \<and> length (flat vA) > length (flat v1a)" using q1 q2 PosOrd_SeqI1 by auto 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1235
     then show "False"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1236
       using PosOrd_shorterI by blast  
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1237
  qed
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1238
  ultimately
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1239
  show "s \<in> SEQ r1 r2 \<rightarrow> v1" unfolding eqs
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1240
    by (rule Posix.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1241
next
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1242
   case (STAR r s v1)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1243
   have IH: "\<And>s v1. \<lbrakk>v1 \<in> CPT r s; \<forall>v2\<in>PT r s. \<not> v2 :\<sqsubset>val v1\<rbrakk> \<Longrightarrow> s \<in> r \<rightarrow> v1" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1244
   have as1: "\<forall>v2\<in>PT (STAR r) s. \<not> v2 :\<sqsubset>val v1" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1245
   have as2: "v1 \<in> CPT (STAR r) s" by fact
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1246
   then obtain 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1247
    vs where eqs:
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1248
        "v1 = Stars vs" "s = flat (Stars vs)"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1249
        "\<forall>v \<in> set vs. v \<in> CPT r (flat v)"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1250
        unfolding CPT_def by (auto elim: CPrf.cases dest!: CPrf_stars)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1251
   have "Stars vs \<in> CPT (STAR r) (flat (Stars vs))" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1252
     using as2 unfolding eqs .
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1253
   moreover
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1254
   have "\<forall>v\<in>set vs. flat v \<in> r \<rightarrow> v" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1255
     proof 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1256
        fix v
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1257
        assume a: "v \<in> set vs"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1258
        then obtain pre post where e: "vs = pre @ [v] @ post"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1259
          by (metis append_Cons append_Nil in_set_conv_decomp_first)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1260
        then have q: "\<forall>v2\<in>PT (STAR r) s. \<not> v2 :\<sqsubset>val Stars (pre @ [v] @ post)" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1261
          using as1 unfolding eqs by simp
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1262
        have "\<forall>v2\<in>PT r (flat v). \<not> v2 :\<sqsubset>val v" unfolding eqs 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1263
          proof (rule ballI, rule notI) 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1264
             fix v2
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1265
             assume w: "v2 :\<sqsubset>val v"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1266
             assume "v2 \<in> PT r (flat v)"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1267
             then have "Stars (pre @ [v2] @ post) \<in> PT (STAR r) s" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1268
                 using as2 unfolding e eqs
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1269
                 apply(auto simp add: CPT_def PT_def intro!: Prf_Stars)[1]
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1270
                 using CPrf_Stars_appendE CPrf_stars Prf_CPrf apply blast
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1271
                 by (meson CPrf_Stars_appendE CPrf_stars Prf_CPrf list.set_intros(2))
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1272
             then have  "\<not> Stars (pre @ [v2] @ post) :\<sqsubset>val Stars (pre @ [v] @ post)"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1273
                using q by simp     
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1274
             with w show "False"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1275
                using PT_def \<open>v2 \<in> PT r (flat v)\<close> append_Cons flat.simps(7) mem_Collect_eq 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1276
                PosOrd_StarsI PosOrd_Stars_appendI by auto
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1277
          qed     
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1278
        with IH
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1279
        show "flat v \<in> r \<rightarrow> v" using a as2 unfolding eqs
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1280
          using eqs(3) by blast
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1281
     qed
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1282
   moreover
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1283
   have "\<not> (\<exists>vs2\<in>PT (STAR r) (flat (Stars vs)). vs2 :\<sqsubset>val Stars vs)" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1284
     proof 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1285
       assume "\<exists>vs2 \<in> PT (STAR r) (flat (Stars vs)). vs2 :\<sqsubset>val Stars vs"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1286
       then obtain vs2 where "\<turnstile> Stars vs2 : STAR r" "flat (Stars vs2) = flat (Stars vs)"
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1287
                             "Stars vs2 :\<sqsubset>val Stars vs" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1288
         unfolding PT_def
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1289
         apply(auto elim: Prf.cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1290
         apply(erule Prf.cases)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1291
         apply(auto intro: Prf.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1292
         apply(drule_tac x="[]" in meta_spec) 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1293
         apply(simp)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1294
         apply(drule meta_mp)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1295
         apply(auto intro: Prf.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1296
         apply(drule_tac x="(v#vsa)" in meta_spec)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1297
         apply(auto intro: Prf.intros)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1298
         done
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1299
       then show "False" using as1 unfolding eqs
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1300
         apply -
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1301
         apply(drule_tac x="Stars vs2" in bspec)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1302
         apply(auto simp add: PT_def)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1303
         done
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1304
     qed
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1305
   ultimately have "flat (Stars vs) \<in> STAR r \<rightarrow> Stars vs" 
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1306
     by (rule PosOrd_Posix_Stars)
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1307
   then show "s \<in> STAR r \<rightarrow> v1" unfolding eqs .
247fc5dd4943 isar proofs
Christian Urban <urbanc@in.tum.de>
parents: 257
diff changeset
  1308
qed
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1309
249
256484ef4be4 updated
Christian Urban <urbanc@in.tum.de>
parents: 248
diff changeset
  1310
unused_thms
256484ef4be4 updated
Christian Urban <urbanc@in.tum.de>
parents: 248
diff changeset
  1311
248
b90ff5abb437 added a proof that Positional ordering is equivalent to direct posix definition
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1312
end