src/HOL/IMP/Abs_Int_ITP/ACom_ITP.thy
author blanchet
Wed, 24 Sep 2014 15:45:55 +0200
changeset 58425 246985c6b20b
parent 58310 91ea607a34d8
permissions -rw-r--r--
simpler proof
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
     2
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
     3
theory ACom_ITP
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
     4
imports "../Com"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
     5
begin
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
     6
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
     7
subsection "Annotated Commands"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
     8
58310
91ea607a34d8 updated news
blanchet
parents: 58249
diff changeset
     9
datatype 'a acom =
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    10
  SKIP 'a                           ("SKIP {_}" 61) |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    11
  Assign vname aexp 'a              ("(_ ::= _/ {_})" [1000, 61, 0] 61) |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 50986
diff changeset
    12
  Seq "('a acom)" "('a acom)"       ("_;;//_"  [60, 61] 60) |
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    13
  If bexp "('a acom)" "('a acom)" 'a
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    14
    ("(IF _/ THEN _/ ELSE _//{_})"  [0, 0, 61, 0] 61) |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    15
  While 'a bexp "('a acom)" 'a
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    16
    ("({_}//WHILE _/ DO (_)//{_})"  [0, 0, 61, 0] 61)
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    17
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    18
fun post :: "'a acom \<Rightarrow>'a" where
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    19
"post (SKIP {P}) = P" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    20
"post (x ::= e {P}) = P" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 50986
diff changeset
    21
"post (c1;; c2) = post c2" |
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    22
"post (IF b THEN c1 ELSE c2 {P}) = P" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    23
"post ({Inv} WHILE b DO c {P}) = P"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    24
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    25
fun strip :: "'a acom \<Rightarrow> com" where
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    26
"strip (SKIP {P}) = com.SKIP" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    27
"strip (x ::= e {P}) = (x ::= e)" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 50986
diff changeset
    28
"strip (c1;;c2) = (strip c1;; strip c2)" |
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    29
"strip (IF b THEN c1 ELSE c2 {P}) = (IF b THEN strip c1 ELSE strip c2)" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    30
"strip ({Inv} WHILE b DO c {P}) = (WHILE b DO strip c)"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    31
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    32
fun anno :: "'a \<Rightarrow> com \<Rightarrow> 'a acom" where
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    33
"anno a com.SKIP = SKIP {a}" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    34
"anno a (x ::= e) = (x ::= e {a})" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 50986
diff changeset
    35
"anno a (c1;;c2) = (anno a c1;; anno a c2)" |
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    36
"anno a (IF b THEN c1 ELSE c2) =
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    37
  (IF b THEN anno a c1 ELSE anno a c2 {a})" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    38
"anno a (WHILE b DO c) =
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    39
  ({a} WHILE b DO anno a c {a})"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    40
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    41
fun annos :: "'a acom \<Rightarrow> 'a list" where
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    42
"annos (SKIP {a}) = [a]" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    43
"annos (x::=e {a}) = [a]" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 50986
diff changeset
    44
"annos (C1;;C2) = annos C1 @ annos C2" |
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    45
"annos (IF b THEN C1 ELSE C2 {a}) = a #  annos C1 @ annos C2" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    46
"annos ({i} WHILE b DO C {a}) = i # a # annos C"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    47
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    48
fun map_acom :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a acom \<Rightarrow> 'b acom" where
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    49
"map_acom f (SKIP {P}) = SKIP {f P}" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    50
"map_acom f (x ::= e {P}) = (x ::= e {f P})" |
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 50986
diff changeset
    51
"map_acom f (c1;;c2) = (map_acom f c1;; map_acom f c2)" |
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    52
"map_acom f (IF b THEN c1 ELSE c2 {P}) =
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    53
  (IF b THEN map_acom f c1 ELSE map_acom f c2 {f P})" |
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    54
"map_acom f ({Inv} WHILE b DO c {P}) =
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    55
  ({f Inv} WHILE b DO map_acom f c {f P})"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    56
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    57
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    58
lemma post_map_acom[simp]: "post(map_acom f c) = f(post c)"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    59
by (induction c) simp_all
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    60
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    61
lemma strip_acom[simp]: "strip (map_acom f c) = strip c"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    62
by (induction c) auto
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    63
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    64
lemma map_acom_SKIP:
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    65
 "map_acom f c = SKIP {S'} \<longleftrightarrow> (\<exists>S. c = SKIP {S} \<and> S' = f S)"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    66
by (cases c) auto
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    67
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    68
lemma map_acom_Assign:
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    69
 "map_acom f c = x ::= e {S'} \<longleftrightarrow> (\<exists>S. c = x::=e {S} \<and> S' = f S)"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    70
by (cases c) auto
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    71
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    72
lemma map_acom_Seq:
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 50986
diff changeset
    73
 "map_acom f c = c1';;c2' \<longleftrightarrow>
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 50986
diff changeset
    74
 (\<exists>c1 c2. c = c1;;c2 \<and> map_acom f c1 = c1' \<and> map_acom f c2 = c2')"
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    75
by (cases c) auto
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    76
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    77
lemma map_acom_If:
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    78
 "map_acom f c = IF b THEN c1' ELSE c2' {S'} \<longleftrightarrow>
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    79
 (\<exists>S c1 c2. c = IF b THEN c1 ELSE c2 {S} \<and> map_acom f c1 = c1' \<and> map_acom f c2 = c2' \<and> S' = f S)"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    80
by (cases c) auto
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    81
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    82
lemma map_acom_While:
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    83
 "map_acom f w = {I'} WHILE b DO c' {P'} \<longleftrightarrow>
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    84
 (\<exists>I P c. w = {I} WHILE b DO c {P} \<and> map_acom f c = c' \<and> I' = f I \<and> P' = f P)"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    85
by (cases w) auto
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    86
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    87
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    88
lemma strip_anno[simp]: "strip (anno a c) = c"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    89
by(induct c) simp_all
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    90
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    91
lemma strip_eq_SKIP:
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    92
  "strip c = com.SKIP \<longleftrightarrow> (EX P. c = SKIP {P})"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    93
by (cases c) simp_all
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    94
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    95
lemma strip_eq_Assign:
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    96
  "strip c = x::=e \<longleftrightarrow> (EX P. c = x::=e {P})"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    97
by (cases c) simp_all
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    98
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
    99
lemma strip_eq_Seq:
52046
bc01725d7918 replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
nipkow
parents: 50986
diff changeset
   100
  "strip c = c1;;c2 \<longleftrightarrow> (EX d1 d2. c = d1;;d2 & strip d1 = c1 & strip d2 = c2)"
49095
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   101
by (cases c) simp_all
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   102
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   103
lemma strip_eq_If:
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   104
  "strip c = IF b THEN c1 ELSE c2 \<longleftrightarrow>
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   105
  (EX d1 d2 P. c = IF b THEN d1 ELSE d2 {P} & strip d1 = c1 & strip d2 = c2)"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   106
by (cases c) simp_all
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   107
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   108
lemma strip_eq_While:
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   109
  "strip c = WHILE b DO c1 \<longleftrightarrow>
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   110
  (EX I d1 P. c = {I} WHILE b DO d1 {P} & strip d1 = c1)"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   111
by (cases c) simp_all
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   112
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   113
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   114
lemma set_annos_anno[simp]: "set (annos (anno a C)) = {a}"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   115
by(induction C)(auto)
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   116
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   117
lemma size_annos_same: "strip C1 = strip C2 \<Longrightarrow> size(annos C1) = size(annos C2)"
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   118
apply(induct C2 arbitrary: C1)
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   119
apply (auto simp: strip_eq_SKIP strip_eq_Assign strip_eq_Seq strip_eq_If strip_eq_While)
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   120
done
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   121
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   122
lemmas size_annos_same2 = eqTrueI[OF size_annos_same]
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   123
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   124
7df19036392e added annotations after condition in if and while
nipkow
parents:
diff changeset
   125
end