src/HOL/IMP/AbsInt0_const.thy
author haftmann
Thu, 08 Sep 2011 11:31:23 +0200
changeset 44839 d19c677eb812
parent 44656 22bbd0d1b943
child 44932 7c93ee993cae
permissions -rw-r--r--
tuned
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
44656
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     2
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     3
theory AbsInt0_const
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     4
imports AbsInt0
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     5
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     6
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     7
subsection "Constant Propagation"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     8
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
     9
datatype cval = Const val | Any
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    10
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    11
fun rep_cval where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    12
"rep_cval (Const n) = {n}" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    13
"rep_cval (Any) = UNIV"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    14
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    15
fun plus_cval where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    16
"plus_cval (Const m) (Const n) = Const(m+n)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    17
"plus_cval _ _ = Any"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    18
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    19
instantiation cval :: SL_top
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    20
begin
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    21
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    22
fun le_cval where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    23
"_ \<sqsubseteq> Any = True" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    24
"Const n \<sqsubseteq> Const m = (n=m)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    25
"Any \<sqsubseteq> Const _ = False"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    26
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    27
fun join_cval where
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    28
"Const m \<squnion> Const n = (if n=m then Const m else Any)" |
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    29
"_ \<squnion> _ = Any"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    30
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    31
definition "Top = Any"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    32
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    33
instance
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    34
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    35
  case goal1 thus ?case by (cases x) simp_all
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    36
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    37
  case goal2 thus ?case by(cases z, cases y, cases x, simp_all)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    38
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    39
  case goal3 thus ?case by(cases x, cases y, simp_all)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    40
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    41
  case goal4 thus ?case by(cases y, cases x, simp_all)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    42
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    43
  case goal5 thus ?case by(cases z, cases y, cases x, simp_all)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    44
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    45
  case goal6 thus ?case by(simp add: Top_cval_def)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    46
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    47
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    48
end
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    49
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    50
interpretation Rep rep_cval
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    51
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    52
  case goal1 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    53
    by(cases a, cases b, simp, simp, cases b, simp, simp)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    54
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    55
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    56
interpretation Val_abs rep_cval Const plus_cval
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    57
proof
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    58
  case goal1 show ?case by simp
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    59
next
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    60
  case goal2 thus ?case
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    61
    by(cases a1, cases a2, simp, simp, cases a2, simp, simp)
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    62
qed
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    63
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    64
interpretation Abs_Int rep_cval Const plus_cval
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    65
defines AI_const is AI
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    66
and aval'_const is aval'
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    67
..
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    68
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    69
text{* Straight line code: *}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    70
definition "test1_const =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    71
 ''y'' ::= N 7;
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    72
 ''z'' ::= Plus (V ''y'') (N 2);
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    73
 ''y'' ::= Plus (V ''x'') (N 0)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    74
value [code] "list (AI_const test1_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    75
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    76
text{* Conditional: *}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    77
definition "test2_const =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    78
 IF Less (N 41) (V ''x'') THEN ''x'' ::= N 5 ELSE ''x'' ::= N 5"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    79
value "list (AI_const test2_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    80
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    81
text{* Conditional, test is ignored: *}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    82
definition "test3_const =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    83
 ''x'' ::= N 42;
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    84
 IF Less (N 41) (V ''x'') THEN ''x'' ::= N 5 ELSE ''x'' ::= N 6"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    85
value "list (AI_const test3_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    86
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    87
text{* While: *}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    88
definition "test4_const =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    89
 ''x'' ::= N 0; WHILE B True DO ''x'' ::= N 0"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    90
value [code] "list (AI_const test4_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    91
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    92
text{* While, test is ignored: *}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    93
definition "test5_const =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    94
 ''x'' ::= N 0; WHILE Less (V ''x'') (N 1) DO ''x'' ::= N 1"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    95
value [code] "list (AI_const test5_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    96
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    97
text{* Iteration is needed: *}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    98
definition "test6_const =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
    99
  ''x'' ::= N 0; ''y'' ::= N 0; ''z'' ::= N 2;
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   100
  WHILE Less (V ''x'') (N 1) DO (''x'' ::= V ''y''; ''y'' ::= V ''z'')"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   101
value [code] "list (AI_const test6_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   102
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   103
text{* More iteration would be needed: *}
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   104
definition "test7_const =
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   105
  ''x'' ::= N 0; ''y'' ::= N 0; ''z'' ::= N 0; ''u'' ::= N 3;
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   106
  WHILE B True DO (''x'' ::= V ''y''; ''y'' ::= V ''z''; ''z'' ::= V ''u'')"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   107
value [code] "list (AI_const test7_const Top)"
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   108
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   109
22bbd0d1b943 Added Abstract Interpretation theories
nipkow
parents:
diff changeset
   110
end