| author | wenzelm | 
| Thu, 09 Oct 2014 13:56:27 +0200 | |
| changeset 58640 | 37f852399a32 | 
| parent 58310 | 91ea607a34d8 | 
| child 61671 | 20d4cd2ceab2 | 
| permissions | -rw-r--r-- | 
| 44656 | 1 | (* Author: Tobias Nipkow *) | 
| 2 | ||
| 45111 | 3 | theory Abs_Int_den0_const | 
| 4 | imports Abs_Int_den0 | |
| 44656 | 5 | begin | 
| 6 | ||
| 7 | subsection "Constant Propagation" | |
| 8 | ||
| 58310 | 9 | datatype cval = Const val | Any | 
| 44656 | 10 | |
| 11 | fun rep_cval where | |
| 12 | "rep_cval (Const n) = {n}" |
 | |
| 13 | "rep_cval (Any) = UNIV" | |
| 14 | ||
| 15 | fun plus_cval where | |
| 16 | "plus_cval (Const m) (Const n) = Const(m+n)" | | |
| 17 | "plus_cval _ _ = Any" | |
| 18 | ||
| 19 | instantiation cval :: SL_top | |
| 20 | begin | |
| 21 | ||
| 22 | fun le_cval where | |
| 23 | "_ \<sqsubseteq> Any = True" | | |
| 24 | "Const n \<sqsubseteq> Const m = (n=m)" | | |
| 25 | "Any \<sqsubseteq> Const _ = False" | |
| 26 | ||
| 27 | fun join_cval where | |
| 28 | "Const m \<squnion> Const n = (if n=m then Const m else Any)" | | |
| 29 | "_ \<squnion> _ = Any" | |
| 30 | ||
| 31 | definition "Top = Any" | |
| 32 | ||
| 33 | instance | |
| 34 | proof | |
| 35 | case goal1 thus ?case by (cases x) simp_all | |
| 36 | next | |
| 37 | case goal2 thus ?case by(cases z, cases y, cases x, simp_all) | |
| 38 | next | |
| 39 | case goal3 thus ?case by(cases x, cases y, simp_all) | |
| 40 | next | |
| 41 | case goal4 thus ?case by(cases y, cases x, simp_all) | |
| 42 | next | |
| 43 | case goal5 thus ?case by(cases z, cases y, cases x, simp_all) | |
| 44 | next | |
| 45 | case goal6 thus ?case by(simp add: Top_cval_def) | |
| 46 | qed | |
| 47 | ||
| 48 | end | |
| 49 | ||
| 55599 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 haftmann parents: 
52046diff
changeset | 50 | permanent_interpretation Rep rep_cval | 
| 44656 | 51 | proof | 
| 52 | case goal1 thus ?case | |
| 53 | by(cases a, cases b, simp, simp, cases b, simp, simp) | |
| 54 | qed | |
| 55 | ||
| 55599 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 haftmann parents: 
52046diff
changeset | 56 | permanent_interpretation Val_abs rep_cval Const plus_cval | 
| 44656 | 57 | proof | 
| 58 | case goal1 show ?case by simp | |
| 59 | next | |
| 60 | case goal2 thus ?case | |
| 61 | by(cases a1, cases a2, simp, simp, cases a2, simp, simp) | |
| 62 | qed | |
| 63 | ||
| 55599 
6535c537b243
aggiornamento for "interpretation with definitions": operate uniformly on theory and locale level under the brand of "permanent interpretation"
 haftmann parents: 
52046diff
changeset | 64 | permanent_interpretation Abs_Int rep_cval Const plus_cval "(iter' 3)" | 
| 55600 
3c7610b8dcfc
more convenient syntax for permanent interpretation
 haftmann parents: 
55599diff
changeset | 65 | defining AI_const = AI | 
| 
3c7610b8dcfc
more convenient syntax for permanent interpretation
 haftmann parents: 
55599diff
changeset | 66 | and aval'_const = aval' | 
| 44944 | 67 | proof qed (auto simp: iter'_pfp_above) | 
| 44656 | 68 | |
| 69 | text{* Straight line code: *}
 | |
| 70 | definition "test1_const = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
45200diff
changeset | 71 | ''y'' ::= N 7;; | 
| 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
45200diff
changeset | 72 | ''z'' ::= Plus (V ''y'') (N 2);; | 
| 44656 | 73 | ''y'' ::= Plus (V ''x'') (N 0)" | 
| 74 | ||
| 75 | text{* Conditional: *}
 | |
| 76 | definition "test2_const = | |
| 77 | IF Less (N 41) (V ''x'') THEN ''x'' ::= N 5 ELSE ''x'' ::= N 5" | |
| 78 | ||
| 79 | text{* Conditional, test is ignored: *}
 | |
| 80 | definition "test3_const = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
45200diff
changeset | 81 | ''x'' ::= N 42;; | 
| 44656 | 82 | IF Less (N 41) (V ''x'') THEN ''x'' ::= N 5 ELSE ''x'' ::= N 6" | 
| 83 | ||
| 84 | text{* While: *}
 | |
| 85 | definition "test4_const = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
45200diff
changeset | 86 | ''x'' ::= N 0;; WHILE Bc True DO ''x'' ::= N 0" | 
| 44656 | 87 | |
| 88 | text{* While, test is ignored: *}
 | |
| 89 | definition "test5_const = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
45200diff
changeset | 90 | ''x'' ::= N 0;; WHILE Less (V ''x'') (N 1) DO ''x'' ::= N 1" | 
| 44656 | 91 | |
| 92 | text{* Iteration is needed: *}
 | |
| 93 | definition "test6_const = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
45200diff
changeset | 94 | ''x'' ::= N 0;; ''y'' ::= N 0;; ''z'' ::= N 2;; | 
| 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
45200diff
changeset | 95 |   WHILE Less (V ''x'') (N 1) DO (''x'' ::= V ''y'';; ''y'' ::= V ''z'')"
 | 
| 44656 | 96 | |
| 97 | text{* More iteration would be needed: *}
 | |
| 98 | definition "test7_const = | |
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
45200diff
changeset | 99 | ''x'' ::= N 0;; ''y'' ::= N 0;; ''z'' ::= N 0;; ''u'' ::= N 3;; | 
| 44932 | 100 | WHILE Less (V ''x'') (N 1) | 
| 52046 
bc01725d7918
replaced `;' by `;;' to disambiguate syntax; unexpected slight increase in build time
 nipkow parents: 
45200diff
changeset | 101 |   DO (''x'' ::= V ''y'';; ''y'' ::= V ''z'';; ''z'' ::= V ''u'')"
 | 
| 44932 | 102 | |
| 56927 | 103 | value "list (AI_const test1_const Top)" | 
| 104 | value "list (AI_const test2_const Top)" | |
| 105 | value "list (AI_const test3_const Top)" | |
| 106 | value "list (AI_const test4_const Top)" | |
| 107 | value "list (AI_const test5_const Top)" | |
| 108 | value "list (AI_const test6_const Top)" | |
| 109 | value "list (AI_const test7_const Top)" | |
| 44656 | 110 | |
| 111 | end |