| author | wenzelm | 
| Sat, 22 Mar 2025 13:54:18 +0100 | |
| changeset 82318 | d0838fb98fc3 | 
| parent 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 43158 | 1 | theory C_like imports Main begin | 
| 2 | ||
| 3 | subsection "A C-like Language" | |
| 4 | ||
| 5 | type_synonym state = "nat \<Rightarrow> nat" | |
| 6 | ||
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69597diff
changeset | 7 | datatype aexp = N nat | Deref aexp (\<open>!\<close>) | Plus aexp aexp | 
| 43158 | 8 | |
| 9 | fun aval :: "aexp \<Rightarrow> state \<Rightarrow> nat" where | |
| 10 | "aval (N n) s = n" | | |
| 11 | "aval (!a) s = s(aval a s)" | | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 12 | "aval (Plus a\<^sub>1 a\<^sub>2) s = aval a\<^sub>1 s + aval a\<^sub>2 s" | 
| 43158 | 13 | |
| 58310 | 14 | datatype bexp = Bc bool | Not bexp | And bexp bexp | Less aexp aexp | 
| 43158 | 15 | |
| 16 | primrec bval :: "bexp \<Rightarrow> state \<Rightarrow> bool" where | |
| 45200 | 17 | "bval (Bc v) _ = v" | | 
| 43158 | 18 | "bval (Not b) s = (\<not> bval b s)" | | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 19 | "bval (And b\<^sub>1 b\<^sub>2) s = (if bval b\<^sub>1 s then bval b\<^sub>2 s else False)" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 20 | "bval (Less a\<^sub>1 a\<^sub>2) s = (aval a\<^sub>1 s < aval a\<^sub>2 s)" | 
| 43158 | 21 | |
| 22 | ||
| 58310 | 23 | datatype | 
| 43158 | 24 | com = SKIP | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69597diff
changeset | 25 | | Assign aexp aexp (\<open>_ ::= _\<close> [61, 61] 61) | 
| 43158 | 26 | | New aexp aexp | 
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69597diff
changeset | 27 | | Seq com com (\<open>_;/ _\<close> [60, 61] 60) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69597diff
changeset | 28 | | If bexp com com (\<open>(IF _/ THEN _/ ELSE _)\<close> [0, 0, 61] 61) | 
| 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69597diff
changeset | 29 | | While bexp com (\<open>(WHILE _/ DO _)\<close> [0, 61] 61) | 
| 43158 | 30 | |
| 31 | inductive | |
| 80914 
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
 wenzelm parents: 
69597diff
changeset | 32 | big_step :: "com \<times> state \<times> nat \<Rightarrow> state \<times> nat \<Rightarrow> bool" (infix \<open>\<Rightarrow>\<close> 55) | 
| 43158 | 33 | where | 
| 34 | Skip: "(SKIP,sn) \<Rightarrow> sn" | | |
| 35 | Assign: "(lhs ::= a,s,n) \<Rightarrow> (s(aval lhs s := aval a s),n)" | | |
| 36 | New: "(New lhs a,s,n) \<Rightarrow> (s(aval lhs s := n), n+aval a s)" | | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 37 | Seq: "\<lbrakk> (c\<^sub>1,sn\<^sub>1) \<Rightarrow> sn\<^sub>2; (c\<^sub>2,sn\<^sub>2) \<Rightarrow> sn\<^sub>3 \<rbrakk> \<Longrightarrow> | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 38 | (c\<^sub>1;c\<^sub>2, sn\<^sub>1) \<Rightarrow> sn\<^sub>3" | | 
| 43158 | 39 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 40 | IfTrue: "\<lbrakk> bval b s; (c\<^sub>1,s,n) \<Rightarrow> tn \<rbrakk> \<Longrightarrow> | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 41 | (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s,n) \<Rightarrow> tn" | | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 42 | IfFalse: "\<lbrakk> \<not>bval b s; (c\<^sub>2,s,n) \<Rightarrow> tn \<rbrakk> \<Longrightarrow> | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 43 | (IF b THEN c\<^sub>1 ELSE c\<^sub>2, s,n) \<Rightarrow> tn" | | 
| 43158 | 44 | |
| 45 | WhileFalse: "\<not>bval b s \<Longrightarrow> (WHILE b DO c,s,n) \<Rightarrow> (s,n)" | | |
| 46 | WhileTrue: | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 47 | "\<lbrakk> bval b s\<^sub>1; (c,s\<^sub>1,n) \<Rightarrow> sn\<^sub>2; (WHILE b DO c, sn\<^sub>2) \<Rightarrow> sn\<^sub>3 \<rbrakk> \<Longrightarrow> | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
47818diff
changeset | 48 | (WHILE b DO c, s\<^sub>1,n) \<Rightarrow> sn\<^sub>3" | 
| 43158 | 49 | |
| 50 | code_pred big_step . | |
| 51 | ||
| 43738 
e40d2eddf2c0
adding a very liberal timeout for values after a test case failed due to the restricted timeout
 bulwahn parents: 
43158diff
changeset | 52 | declare [[values_timeout = 3600]] | 
| 43158 | 53 | |
| 67406 | 54 | text\<open>Examples:\<close> | 
| 43158 | 55 | |
| 56 | definition | |
| 57 | "array_sum = | |
| 58 | WHILE Less (!(N 0)) (Plus (!(N 1)) (N 1)) | |
| 59 | DO ( N 2 ::= Plus (!(N 2)) (!(!(N 0))); | |
| 60 | N 0 ::= Plus (!(N 0)) (N 1) )" | |
| 61 | ||
| 69597 | 62 | text \<open>To show the first n variables in a \<^typ>\<open>nat \<Rightarrow> nat\<close> state:\<close> | 
| 43158 | 63 | definition | 
| 64 | "list t n = map t [0 ..< n]" | |
| 65 | ||
| 66 | values "{list t n |t n. (array_sum, nth[3,4,0,3,7],5) \<Rightarrow> (t,n)}"
 | |
| 67 | ||
| 68 | definition | |
| 69 | "linked_list_sum = | |
| 70 | WHILE Less (N 0) (!(N 0)) | |
| 71 | DO ( N 1 ::= Plus(!(N 1)) (!(!(N 0))); | |
| 72 | N 0 ::= !(Plus(!(N 0))(N 1)) )" | |
| 73 | ||
| 74 | values "{list t n |t n. (linked_list_sum, nth[4,0,3,0,7,2],6) \<Rightarrow> (t,n)}"
 | |
| 75 | ||
| 76 | definition | |
| 77 | "array_init = | |
| 78 | New (N 0) (!(N 1)); N 2 ::= !(N 0); | |
| 79 | WHILE Less (!(N 2)) (Plus (!(N 0)) (!(N 1))) | |
| 80 | DO ( !(N 2) ::= !(N 2); | |
| 81 | N 2 ::= Plus (!(N 2)) (N 1) )" | |
| 82 | ||
| 83 | values "{list t n |t n. (array_init, nth[5,2,7],3) \<Rightarrow> (t,n)}"
 | |
| 84 | ||
| 85 | definition | |
| 86 | "linked_list_init = | |
| 87 | WHILE Less (!(N 1)) (!(N 0)) | |
| 88 | DO ( New (N 3) (N 2); | |
| 89 | N 1 ::= Plus (!(N 1)) (N 1); | |
| 90 | !(N 3) ::= !(N 1); | |
| 91 | Plus (!(N 3)) (N 1) ::= !(N 2); | |
| 92 | N 2 ::= !(N 3) )" | |
| 93 | ||
| 94 | values "{list t n |t n. (linked_list_init, nth[2,0,0,0],4) \<Rightarrow> (t,n)}"
 | |
| 95 | ||
| 96 | end |