src/HOL/IMP/Compiler.thy
author haftmann
Fri, 16 Mar 2007 21:32:10 +0100
changeset 22454 c3654ba76a09
parent 20217 25b068a99d2b
child 23746 a455e69c31cc
permissions -rw-r--r--
integrated with LOrder.thy
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10343
24c87e1366d8 *** empty log message ***
nipkow
parents: 10342
diff changeset
     1
(*  Title:      HOL/IMP/Compiler.thy
24c87e1366d8 *** empty log message ***
nipkow
parents: 10342
diff changeset
     2
    ID:         $Id$
24c87e1366d8 *** empty log message ***
nipkow
parents: 10342
diff changeset
     3
    Author:     Tobias Nipkow, TUM
24c87e1366d8 *** empty log message ***
nipkow
parents: 10342
diff changeset
     4
    Copyright   1996 TUM
12429
15c13bdc94c8 tuned for latex output
kleing
parents: 12275
diff changeset
     5
*)
10343
24c87e1366d8 *** empty log message ***
nipkow
parents: 10342
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14738
diff changeset
     7
theory Compiler imports Machines begin
12429
15c13bdc94c8 tuned for latex output
kleing
parents: 12275
diff changeset
     8
15c13bdc94c8 tuned for latex output
kleing
parents: 12275
diff changeset
     9
subsection "The compiler"
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
    10
12429
15c13bdc94c8 tuned for latex output
kleing
parents: 12275
diff changeset
    11
consts compile :: "com \<Rightarrow> instr list"
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
    12
primrec
12429
15c13bdc94c8 tuned for latex output
kleing
parents: 12275
diff changeset
    13
"compile \<SKIP> = []"
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13630
diff changeset
    14
"compile (x:==a) = [SET x a]"
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
    15
"compile (c1;c2) = compile c1 @ compile c2"
12429
15c13bdc94c8 tuned for latex output
kleing
parents: 12275
diff changeset
    16
"compile (\<IF> b \<THEN> c1 \<ELSE> c2) =
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    17
 [JMPF b (length(compile c1) + 1)] @ compile c1 @
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13630
diff changeset
    18
 [JMPF (\<lambda>x. False) (length(compile c2))] @ compile c2"
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    19
"compile (\<WHILE> b \<DO> c) = [JMPF b (length(compile c) + 1)] @ compile c @
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
    20
 [JMPB (length(compile c)+1)]"
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
    21
12429
15c13bdc94c8 tuned for latex output
kleing
parents: 12275
diff changeset
    22
subsection "Compiler correctness"
11275
71498de45241 new proof
nipkow
parents: 10425
diff changeset
    23
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    24
theorem assumes A: "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    25
shows "\<And>p q. \<langle>compile c @ p,q,s\<rangle> -*\<rightarrow> \<langle>p,rev(compile c)@q,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    26
  (is "\<And>p q. ?P c s t p q")
12275
aa2b7b475a94 Isar conversion
nipkow
parents: 11704
diff changeset
    27
proof -
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    28
  from A show "\<And>p q. ?thesis p q"
12275
aa2b7b475a94 Isar conversion
nipkow
parents: 11704
diff changeset
    29
  proof induct
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    30
    case Skip thus ?case by simp
12275
aa2b7b475a94 Isar conversion
nipkow
parents: 11704
diff changeset
    31
  next
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    32
    case Assign thus ?case by force
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    33
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    34
    case Semi thus ?case by simp (blast intro:rtrancl_trans)
12275
aa2b7b475a94 Isar conversion
nipkow
parents: 11704
diff changeset
    35
  next
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    36
    fix b c0 c1 s0 s1 p q
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    37
    assume IH: "\<And>p q. ?P c0 s0 s1 p q"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    38
    assume "b s0"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    39
    thus "?P (\<IF> b \<THEN> c0 \<ELSE> c1) s0 s1 p q"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    40
      by(simp add: IH[THEN rtrancl_trans])
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    41
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    42
    case IfFalse thus ?case by(simp)
12275
aa2b7b475a94 Isar conversion
nipkow
parents: 11704
diff changeset
    43
  next
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    44
    case WhileFalse thus ?case by simp
12275
aa2b7b475a94 Isar conversion
nipkow
parents: 11704
diff changeset
    45
  next
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    46
    fix b c and s0::state and s1 s2 p q
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    47
    assume b: "b s0" and
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    48
      IHc: "\<And>p q. ?P c s0 s1 p q" and
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    49
      IHw: "\<And>p q. ?P (\<WHILE> b \<DO> c) s1 s2 p q"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    50
    show "?P (\<WHILE> b \<DO> c) s0 s2 p q"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    51
      using b  IHc[THEN rtrancl_trans] IHw by(simp)
12275
aa2b7b475a94 Isar conversion
nipkow
parents: 11704
diff changeset
    52
  qed
aa2b7b475a94 Isar conversion
nipkow
parents: 11704
diff changeset
    53
qed
11275
71498de45241 new proof
nipkow
parents: 10425
diff changeset
    54
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    55
text {* The other direction! *}
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    56
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    57
inductive_cases [elim!]: "(([],p,s),next) : stepa1"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    58
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    59
lemma [simp]: "(\<langle>[],q,s\<rangle> -n\<rightarrow> \<langle>p',q',t\<rangle>) = (n=0 \<and> p' = [] \<and> q' = q \<and> t = s)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    60
apply(rule iffI)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    61
 apply(erule converse_rel_powE, simp, fast)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    62
apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    63
done
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    64
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    65
lemma [simp]: "(\<langle>[],q,s\<rangle> -*\<rightarrow> \<langle>p',q',t\<rangle>) = (p' = [] \<and> q' = q \<and> t = s)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    66
by(simp add: rtrancl_is_UN_rel_pow)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    67
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    68
constdefs
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    69
 forws :: "instr \<Rightarrow> nat set"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    70
"forws instr == case instr of
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13630
diff changeset
    71
 SET x a \<Rightarrow> {0} |
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    72
 JMPF b n \<Rightarrow> {0,n} |
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    73
 JMPB n \<Rightarrow> {}"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    74
 backws :: "instr \<Rightarrow> nat set"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    75
"backws instr == case instr of
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13630
diff changeset
    76
 SET x a \<Rightarrow> {} |
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    77
 JMPF b n \<Rightarrow> {} |
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    78
 JMPB n \<Rightarrow> {n}"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    79
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    80
consts closed :: "nat \<Rightarrow> nat \<Rightarrow> instr list \<Rightarrow> bool"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    81
primrec
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    82
"closed m n [] = True"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    83
"closed m n (instr#is) = ((\<forall>j \<in> forws instr. j \<le> size is+n) \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    84
                        (\<forall>j \<in> backws instr. j \<le> m) \<and> closed (Suc m) n is)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    85
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    86
lemma [simp]:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    87
 "\<And>m n. closed m n (C1@C2) =
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    88
         (closed m (n+size C2) C1 \<and> closed (m+size C1) n C2)"
14738
83f1a514dcb4 changes made due to new Ring_and_Field theory
obua
parents: 13675
diff changeset
    89
by(induct C1, simp, simp add:add_ac)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    90
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    91
theorem [simp]: "\<And>m n. closed m n (compile c)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    92
by(induct c, simp_all add:backws_def forws_def)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    93
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    94
lemma drop_lem: "n \<le> size(p1@p2)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    95
 \<Longrightarrow> (p1' @ p2 = drop n p1 @ drop (n - size p1) p2) =
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    96
    (n \<le> size p1 & p1' = drop n p1)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    97
apply(rule iffI)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    98
 defer apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
    99
apply(subgoal_tac "n \<le> size p1")
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   100
 apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   101
apply(rule ccontr)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   102
apply(drule_tac f = length in arg_cong)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   103
apply simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   104
done
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   105
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   106
lemma reduce_exec1:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   107
 "\<langle>i # p1 @ p2,q1 @ q2,s\<rangle> -1\<rightarrow> \<langle>p1' @ p2,q1' @ q2,s'\<rangle> \<Longrightarrow>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   108
  \<langle>i # p1,q1,s\<rangle> -1\<rightarrow> \<langle>p1',q1',s'\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   109
by(clarsimp simp add: drop_lem split:instr.split_asm split_if_asm)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   110
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   111
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   112
lemma closed_exec1:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   113
 "\<lbrakk> closed 0 0 (rev q1 @ instr # p1);
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   114
    \<langle>instr # p1 @ p2, q1 @ q2,r\<rangle> -1\<rightarrow> \<langle>p',q',r'\<rangle> \<rbrakk> \<Longrightarrow>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   115
  \<exists>p1' q1'. p' = p1'@p2 \<and> q' = q1'@q2 \<and> rev q1' @ p1' = rev q1 @ instr # p1"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   116
apply(clarsimp simp add:forws_def backws_def
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   117
               split:instr.split_asm split_if_asm)
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
   118
done
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
   119
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   120
theorem closed_execn_decomp: "\<And>C1 C2 r.
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   121
 \<lbrakk> closed 0 0 (rev C1 @ C2);
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   122
   \<langle>C2 @ p1 @ p2, C1 @ q,r\<rangle> -n\<rightarrow> \<langle>p2,rev p1 @ rev C2 @ C1 @ q,t\<rangle> \<rbrakk>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   123
 \<Longrightarrow> \<exists>s n1 n2. \<langle>C2,C1,r\<rangle> -n1\<rightarrow> \<langle>[],rev C2 @ C1,s\<rangle> \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   124
     \<langle>p1@p2,rev C2 @ C1 @ q,s\<rangle> -n2\<rightarrow> \<langle>p2, rev p1 @ rev C2 @ C1 @ q,t\<rangle> \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   125
         n = n1+n2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   126
(is "\<And>C1 C2 r. \<lbrakk>?CL C1 C2; ?H C1 C2 r n\<rbrakk> \<Longrightarrow> ?P C1 C2 r n")
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   127
proof(induct n)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   128
  fix C1 C2 r
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   129
  assume "?H C1 C2 r 0"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   130
  thus "?P C1 C2 r 0" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   131
next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   132
  fix C1 C2 r n
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   133
  assume IH: "\<And>C1 C2 r. ?CL C1 C2 \<Longrightarrow> ?H C1 C2 r n \<Longrightarrow> ?P C1 C2 r n"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   134
  assume CL: "?CL C1 C2" and H: "?H C1 C2 r (Suc n)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   135
  show "?P C1 C2 r (Suc n)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   136
  proof (cases C2)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   137
    assume "C2 = []" with H show ?thesis by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   138
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   139
    fix instr tlC2
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   140
    assume C2: "C2 = instr # tlC2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   141
    from H C2 obtain p' q' r'
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   142
      where 1: "\<langle>instr # tlC2 @ p1 @ p2, C1 @ q,r\<rangle> -1\<rightarrow> \<langle>p',q',r'\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   143
      and n: "\<langle>p',q',r'\<rangle> -n\<rightarrow> \<langle>p2,rev p1 @ rev C2 @ C1 @ q,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   144
      by(fastsimp simp add:R_O_Rn_commute)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   145
    from CL closed_exec1[OF _ 1] C2
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   146
    obtain C2' C1' where pq': "p' = C2' @ p1 @ p2 \<and> q' = C1' @ q"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   147
      and same: "rev C1' @ C2' = rev C1 @ C2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   148
      by fastsimp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   149
    have rev_same: "rev C2' @ C1' = rev C2 @ C1"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   150
    proof -
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   151
      have "rev C2' @ C1' = rev(rev C1' @ C2')" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   152
      also have "\<dots> = rev(rev C1 @ C2)" by(simp only:same)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   153
      also have "\<dots> =  rev C2 @ C1" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   154
      finally show ?thesis .
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   155
    qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   156
    hence rev_same': "\<And>p. rev C2' @ C1' @ p = rev C2 @ C1 @ p" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   157
    from n have n': "\<langle>C2' @ p1 @ p2,C1' @ q,r'\<rangle> -n\<rightarrow>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   158
                     \<langle>p2,rev p1 @ rev C2' @ C1' @ q,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   159
      by(simp add:pq' rev_same')
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   160
    from IH[OF _ n'] CL
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   161
    obtain s n1 n2 where n1: "\<langle>C2',C1',r'\<rangle> -n1\<rightarrow> \<langle>[],rev C2 @ C1,s\<rangle>" and
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   162
      "\<langle>p1 @ p2,rev C2 @ C1 @ q,s\<rangle> -n2\<rightarrow> \<langle>p2,rev p1 @ rev C2 @ C1 @ q,t\<rangle> \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   163
       n = n1 + n2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   164
      by(fastsimp simp add: same rev_same rev_same')
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   165
    moreover
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   166
    from 1 n1 pq' C2 have "\<langle>C2,C1,r\<rangle> -Suc n1\<rightarrow> \<langle>[],rev C2 @ C1,s\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   167
      by (simp del:relpow.simps exec_simp) (fast dest:reduce_exec1)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   168
    ultimately show ?thesis by (fastsimp simp del:relpow.simps)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   169
  qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   170
qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   171
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   172
lemma execn_decomp:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   173
"\<langle>compile c @ p1 @ p2,q,r\<rangle> -n\<rightarrow> \<langle>p2,rev p1 @ rev(compile c) @ q,t\<rangle>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   174
 \<Longrightarrow> \<exists>s n1 n2. \<langle>compile c,[],r\<rangle> -n1\<rightarrow> \<langle>[],rev(compile c),s\<rangle> \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   175
     \<langle>p1@p2,rev(compile c) @ q,s\<rangle> -n2\<rightarrow> \<langle>p2, rev p1 @ rev(compile c) @ q,t\<rangle> \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   176
         n = n1+n2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   177
using closed_execn_decomp[of "[]",simplified] by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   178
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   179
lemma exec_star_decomp:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   180
"\<langle>compile c @ p1 @ p2,q,r\<rangle> -*\<rightarrow> \<langle>p2,rev p1 @ rev(compile c) @ q,t\<rangle>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   181
 \<Longrightarrow> \<exists>s. \<langle>compile c,[],r\<rangle> -*\<rightarrow> \<langle>[],rev(compile c),s\<rangle> \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   182
     \<langle>p1@p2,rev(compile c) @ q,s\<rangle> -*\<rightarrow> \<langle>p2, rev p1 @ rev(compile c) @ q,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   183
by(simp add:rtrancl_is_UN_rel_pow)(fast dest: execn_decomp)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   184
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   185
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   186
(* Alternative:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   187
lemma exec_comp_n:
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   188
"\<And>p1 p2 q r t n.
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   189
 \<langle>compile c @ p1 @ p2,q,r\<rangle> -n\<rightarrow> \<langle>p2,rev p1 @ rev(compile c) @ q,t\<rangle>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   190
 \<Longrightarrow> \<exists>s n1 n2. \<langle>compile c,[],r\<rangle> -n1\<rightarrow> \<langle>[],rev(compile c),s\<rangle> \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   191
     \<langle>p1@p2,rev(compile c) @ q,s\<rangle> -n2\<rightarrow> \<langle>p2, rev p1 @ rev(compile c) @ q,t\<rangle> \<and>
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   192
         n = n1+n2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   193
 (is "\<And>p1 p2 q r t n. ?H c p1 p2 q r t n \<Longrightarrow> ?P c p1 p2 q r t n")
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   194
proof (induct c)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   195
*)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   196
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   197
text{*Warning: 
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   198
@{prop"\<langle>compile c @ p,q,s\<rangle> -*\<rightarrow> \<langle>p,rev(compile c)@q,t\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"}
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   199
is not true! *}
10342
b124d59f7b61 *** empty log message ***
nipkow
parents:
diff changeset
   200
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   201
theorem "\<And>s t.
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   202
 \<langle>compile c,[],s\<rangle> -*\<rightarrow> \<langle>[],rev(compile c),t\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   203
proof (induct c)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   204
  fix s t
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   205
  assume "\<langle>compile SKIP,[],s\<rangle> -*\<rightarrow> \<langle>[],rev(compile SKIP),t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   206
  thus "\<langle>SKIP,s\<rangle> \<longrightarrow>\<^sub>c t" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   207
next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   208
  fix s t v f
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   209
  assume "\<langle>compile(v :== f),[],s\<rangle> -*\<rightarrow> \<langle>[],rev(compile(v :== f)),t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   210
  thus "\<langle>v :== f,s\<rangle> \<longrightarrow>\<^sub>c t" by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   211
next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   212
  fix s1 s3 c1 c2
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   213
  let ?C1 = "compile c1" let ?C2 = "compile c2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   214
  assume IH1: "\<And>s t. \<langle>?C1,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C1,t\<rangle> \<Longrightarrow> \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   215
     and IH2: "\<And>s t. \<langle>?C2,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C2,t\<rangle> \<Longrightarrow> \<langle>c2,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   216
  assume "\<langle>compile(c1;c2),[],s1\<rangle> -*\<rightarrow> \<langle>[],rev(compile(c1;c2)),s3\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   217
  then obtain s2 where exec1: "\<langle>?C1,[],s1\<rangle> -*\<rightarrow> \<langle>[],rev ?C1,s2\<rangle>" and
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   218
             exec2: "\<langle>?C2,rev ?C1,s2\<rangle> -*\<rightarrow> \<langle>[],rev(compile(c1;c2)),s3\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   219
    by(fastsimp dest:exec_star_decomp[of _ _ "[]" "[]",simplified])
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   220
  from exec2 have exec2': "\<langle>?C2,[],s2\<rangle> -*\<rightarrow> \<langle>[],rev ?C2,s3\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   221
    using exec_star_decomp[of _ "[]" "[]"] by fastsimp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   222
  have "\<langle>c1,s1\<rangle> \<longrightarrow>\<^sub>c s2" using IH1 exec1 by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   223
  moreover have "\<langle>c2,s2\<rangle> \<longrightarrow>\<^sub>c s3" using IH2 exec2' by fastsimp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   224
  ultimately show "\<langle>c1;c2,s1\<rangle> \<longrightarrow>\<^sub>c s3" ..
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   225
next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   226
  fix s t b c1 c2
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   227
  let ?if = "IF b THEN c1 ELSE c2" let ?C = "compile ?if"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   228
  let ?C1 = "compile c1" let ?C2 = "compile c2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   229
  assume IH1: "\<And>s t. \<langle>?C1,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C1,t\<rangle> \<Longrightarrow> \<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   230
     and IH2: "\<And>s t. \<langle>?C2,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C2,t\<rangle> \<Longrightarrow> \<langle>c2,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   231
     and H: "\<langle>?C,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   232
  show "\<langle>?if,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   233
  proof cases
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   234
    assume b: "b s"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   235
    with H have "\<langle>?C1,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C1,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   236
      by (fastsimp dest:exec_star_decomp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   237
            [of _ "[JMPF (\<lambda>x. False) (size ?C2)]@?C2" "[]",simplified])
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   238
    hence "\<langle>c1,s\<rangle> \<longrightarrow>\<^sub>c t" by(rule IH1)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   239
    with b show ?thesis ..
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   240
  next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   241
    assume b: "\<not> b s"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   242
    with H have "\<langle>?C2,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C2,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   243
      using exec_star_decomp[of _ "[]" "[]"] by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   244
    hence "\<langle>c2,s\<rangle> \<longrightarrow>\<^sub>c t" by(rule IH2)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   245
    with b show ?thesis ..
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   246
  qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   247
next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   248
  fix b c s t
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   249
  let ?w = "WHILE b DO c" let ?W = "compile ?w" let ?C = "compile c"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   250
  let ?j1 = "JMPF b (size ?C + 1)" let ?j2 = "JMPB (size ?C + 1)"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   251
  assume IHc: "\<And>s t. \<langle>?C,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C,t\<rangle> \<Longrightarrow> \<langle>c,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   252
     and H: "\<langle>?W,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?W,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   253
  from H obtain k where ob:"\<langle>?W,[],s\<rangle> -k\<rightarrow> \<langle>[],rev ?W,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   254
    by(simp add:rtrancl_is_UN_rel_pow) blast
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   255
  { fix n have "\<And>s. \<langle>?W,[],s\<rangle> -n\<rightarrow> \<langle>[],rev ?W,t\<rangle> \<Longrightarrow> \<langle>?w,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   256
    proof (induct n rule: less_induct)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   257
      fix n
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   258
      assume IHm: "\<And>m s. \<lbrakk>m < n; \<langle>?W,[],s\<rangle> -m\<rightarrow> \<langle>[],rev ?W,t\<rangle> \<rbrakk> \<Longrightarrow> \<langle>?w,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   259
      fix s
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   260
      assume H: "\<langle>?W,[],s\<rangle> -n\<rightarrow> \<langle>[],rev ?W,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   261
      show "\<langle>?w,s\<rangle> \<longrightarrow>\<^sub>c t"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   262
      proof cases
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   263
	assume b: "b s"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   264
	then obtain m where m: "n = Suc m"
13675
01fc1fc61384 ASIN -> SET
nipkow
parents: 13630
diff changeset
   265
          and "\<langle>?C @ [?j2],[?j1],s\<rangle> -m\<rightarrow> \<langle>[],rev ?W,t\<rangle>"
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   266
	  using H by fastsimp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   267
	then obtain r n1 n2 where n1: "\<langle>?C,[],s\<rangle> -n1\<rightarrow> \<langle>[],rev ?C,r\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   268
          and n2: "\<langle>[?j2],rev ?C @ [?j1],r\<rangle> -n2\<rightarrow> \<langle>[],rev ?W,t\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   269
          and n12: "m = n1+n2"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   270
	  using execn_decomp[of _ "[?j2]"]
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   271
	  by(simp del: execn_simp) fast
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   272
	have n2n: "n2 - 1 < n" using m n12 by arith
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   273
	note b
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   274
	moreover
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   275
	{ from n1 have "\<langle>?C,[],s\<rangle> -*\<rightarrow> \<langle>[],rev ?C,r\<rangle>"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   276
	    by (simp add:rtrancl_is_UN_rel_pow) fast
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   277
	  hence "\<langle>c,s\<rangle> \<longrightarrow>\<^sub>c r" by(rule IHc)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   278
	}
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   279
	moreover
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   280
	{ have "n2 - 1 < n" using m n12 by arith
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   281
	  moreover from n2 have "\<langle>?W,[],r\<rangle> -n2- 1\<rightarrow> \<langle>[],rev ?W,t\<rangle>" by fastsimp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   282
	  ultimately have "\<langle>?w,r\<rangle> \<longrightarrow>\<^sub>c t" by(rule IHm)
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   283
	}
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   284
	ultimately show ?thesis ..
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   285
      next
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   286
	assume b: "\<not> b s"
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   287
	hence "t = s" using H by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   288
	with b show ?thesis by simp
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   289
      qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   290
    qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   291
  }
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   292
  with ob show "\<langle>?w,s\<rangle> \<longrightarrow>\<^sub>c t" by fast
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   293
qed
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   294
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 16417
diff changeset
   295
(* TODO: connect with Machine 0 using M_equiv *)
13095
8ed413a57bdc New machine architecture and other direction of compiler proof.
nipkow
parents: 12566
diff changeset
   296
20217
25b068a99d2b linear arithmetic splits certain operators (e.g. min, max, abs)
webertj
parents: 16417
diff changeset
   297
end