src/HOL/Predicate_Compile_Examples/Examples.thy
changeset 39655 8ad7fe9d6f0b
child 41413 64cd30d6b0b8
equal deleted inserted replaced
39654:1207e39f0c7f 39655:8ad7fe9d6f0b
       
     1 theory Examples
       
     2 imports Main Predicate_Compile_Alternative_Defs
       
     3 begin
       
     4 
       
     5 section {* Formal Languages *}
       
     6 
       
     7 subsection {* General Context Free Grammars *}
       
     8 
       
     9 text {* a contribution by Aditi Barthwal *}
       
    10 
       
    11 datatype ('nts,'ts) symbol = NTS 'nts
       
    12                             | TS 'ts
       
    13 
       
    14                             
       
    15 datatype ('nts,'ts) rule = rule 'nts "('nts,'ts) symbol list"
       
    16 
       
    17 types ('nts,'ts) grammar = "('nts,'ts) rule set * 'nts"
       
    18 
       
    19 fun rules :: "('nts,'ts) grammar => ('nts,'ts) rule set"
       
    20 where
       
    21   "rules (r, s) = r"
       
    22 
       
    23 definition derives 
       
    24 where
       
    25 "derives g = { (lsl,rsl). \<exists>s1 s2 lhs rhs. 
       
    26                          (s1 @ [NTS lhs] @ s2 = lsl) \<and>
       
    27                          (s1 @ rhs @ s2) = rsl \<and>
       
    28                          (rule lhs rhs) \<in> fst g }"
       
    29 
       
    30 abbreviation "example_grammar == 
       
    31 ({ rule ''S'' [NTS ''A'', NTS ''B''],
       
    32    rule ''S'' [TS ''a''],
       
    33   rule ''A'' [TS ''b'']}, ''S'')"
       
    34 
       
    35 
       
    36 code_pred [inductify, skip_proof] derives .
       
    37 
       
    38 thm derives.equation
       
    39 
       
    40 definition "test = { rhs. derives example_grammar ([NTS ''S''], rhs) }"
       
    41 
       
    42 code_pred (modes: o \<Rightarrow> bool) [inductify] test .
       
    43 thm test.equation
       
    44 
       
    45 values "{rhs. test rhs}"
       
    46 
       
    47 declare rtrancl.intros(1)[code_pred_def] converse_rtrancl_into_rtrancl[code_pred_def]
       
    48 
       
    49 code_pred [inductify] rtrancl .
       
    50 
       
    51 definition "test2 = { rhs. ([NTS ''S''],rhs) \<in> (derives example_grammar)^*  }"
       
    52 
       
    53 code_pred [inductify, skip_proof] test2 .
       
    54 
       
    55 values "{rhs. test2 rhs}"
       
    56 
       
    57 subsection {* Some concrete Context Free Grammars *}
       
    58 
       
    59 datatype alphabet = a | b
       
    60 
       
    61 inductive_set S\<^isub>1 and A\<^isub>1 and B\<^isub>1 where
       
    62   "[] \<in> S\<^isub>1"
       
    63 | "w \<in> A\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
       
    64 | "w \<in> B\<^isub>1 \<Longrightarrow> a # w \<in> S\<^isub>1"
       
    65 | "w \<in> S\<^isub>1 \<Longrightarrow> a # w \<in> A\<^isub>1"
       
    66 | "w \<in> S\<^isub>1 \<Longrightarrow> b # w \<in> S\<^isub>1"
       
    67 | "\<lbrakk>v \<in> B\<^isub>1; v \<in> B\<^isub>1\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>1"
       
    68 
       
    69 code_pred [inductify] S\<^isub>1p .
       
    70 code_pred [random_dseq inductify] S\<^isub>1p .
       
    71 thm S\<^isub>1p.equation
       
    72 thm S\<^isub>1p.random_dseq_equation
       
    73 
       
    74 values [random_dseq 5, 5, 5] 5 "{x. S\<^isub>1p x}"
       
    75 
       
    76 inductive_set S\<^isub>2 and A\<^isub>2 and B\<^isub>2 where
       
    77   "[] \<in> S\<^isub>2"
       
    78 | "w \<in> A\<^isub>2 \<Longrightarrow> b # w \<in> S\<^isub>2"
       
    79 | "w \<in> B\<^isub>2 \<Longrightarrow> a # w \<in> S\<^isub>2"
       
    80 | "w \<in> S\<^isub>2 \<Longrightarrow> a # w \<in> A\<^isub>2"
       
    81 | "w \<in> S\<^isub>2 \<Longrightarrow> b # w \<in> B\<^isub>2"
       
    82 | "\<lbrakk>v \<in> B\<^isub>2; v \<in> B\<^isub>2\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>2"
       
    83 
       
    84 code_pred [random_dseq inductify] S\<^isub>2p .
       
    85 thm S\<^isub>2p.random_dseq_equation
       
    86 thm A\<^isub>2p.random_dseq_equation
       
    87 thm B\<^isub>2p.random_dseq_equation
       
    88 
       
    89 values [random_dseq 5, 5, 5] 10 "{x. S\<^isub>2p x}"
       
    90 
       
    91 inductive_set S\<^isub>3 and A\<^isub>3 and B\<^isub>3 where
       
    92   "[] \<in> S\<^isub>3"
       
    93 | "w \<in> A\<^isub>3 \<Longrightarrow> b # w \<in> S\<^isub>3"
       
    94 | "w \<in> B\<^isub>3 \<Longrightarrow> a # w \<in> S\<^isub>3"
       
    95 | "w \<in> S\<^isub>3 \<Longrightarrow> a # w \<in> A\<^isub>3"
       
    96 | "w \<in> S\<^isub>3 \<Longrightarrow> b # w \<in> B\<^isub>3"
       
    97 | "\<lbrakk>v \<in> B\<^isub>3; w \<in> B\<^isub>3\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>3"
       
    98 
       
    99 code_pred [inductify, skip_proof] S\<^isub>3p .
       
   100 thm S\<^isub>3p.equation
       
   101 
       
   102 values 10 "{x. S\<^isub>3p x}"
       
   103 
       
   104 inductive_set S\<^isub>4 and A\<^isub>4 and B\<^isub>4 where
       
   105   "[] \<in> S\<^isub>4"
       
   106 | "w \<in> A\<^isub>4 \<Longrightarrow> b # w \<in> S\<^isub>4"
       
   107 | "w \<in> B\<^isub>4 \<Longrightarrow> a # w \<in> S\<^isub>4"
       
   108 | "w \<in> S\<^isub>4 \<Longrightarrow> a # w \<in> A\<^isub>4"
       
   109 | "\<lbrakk>v \<in> A\<^isub>4; w \<in> A\<^isub>4\<rbrakk> \<Longrightarrow> b # v @ w \<in> A\<^isub>4"
       
   110 | "w \<in> S\<^isub>4 \<Longrightarrow> b # w \<in> B\<^isub>4"
       
   111 | "\<lbrakk>v \<in> B\<^isub>4; w \<in> B\<^isub>4\<rbrakk> \<Longrightarrow> a # v @ w \<in> B\<^isub>4"
       
   112 
       
   113 code_pred (expected_modes: o => bool, i => bool) S\<^isub>4p .
       
   114 
       
   115 hide_const a b
       
   116 
       
   117 section {* Semantics of programming languages *}
       
   118 
       
   119 subsection {* IMP *}
       
   120 
       
   121 types
       
   122   var = nat
       
   123   state = "int list"
       
   124 
       
   125 datatype com =
       
   126   Skip |
       
   127   Ass var "state => int" |
       
   128   Seq com com |
       
   129   IF "state => bool" com com |
       
   130   While "state => bool" com
       
   131 
       
   132 inductive exec :: "com => state => state => bool" where
       
   133 "exec Skip s s" |
       
   134 "exec (Ass x e) s (s[x := e(s)])" |
       
   135 "exec c1 s1 s2 ==> exec c2 s2 s3 ==> exec (Seq c1 c2) s1 s3" |
       
   136 "b s ==> exec c1 s t ==> exec (IF b c1 c2) s t" |
       
   137 "~b s ==> exec c2 s t ==> exec (IF b c1 c2) s t" |
       
   138 "~b s ==> exec (While b c) s s" |
       
   139 "b s1 ==> exec c s1 s2 ==> exec (While b c) s2 s3 ==> exec (While b c) s1 s3"
       
   140 
       
   141 code_pred exec .
       
   142 
       
   143 values "{t. exec
       
   144  (While (%s. s!0 > 0) (Seq (Ass 0 (%s. s!0 - 1)) (Ass 1 (%s. s!1 + 1))))
       
   145  [3,5] t}"
       
   146 
       
   147 subsection {* Lambda *}
       
   148 
       
   149 datatype type =
       
   150     Atom nat
       
   151   | Fun type type    (infixr "\<Rightarrow>" 200)
       
   152 
       
   153 datatype dB =
       
   154     Var nat
       
   155   | App dB dB (infixl "\<degree>" 200)
       
   156   | Abs type dB
       
   157 
       
   158 primrec
       
   159   nth_el :: "'a list \<Rightarrow> nat \<Rightarrow> 'a option" ("_\<langle>_\<rangle>" [90, 0] 91)
       
   160 where
       
   161   "[]\<langle>i\<rangle> = None"
       
   162 | "(x # xs)\<langle>i\<rangle> = (case i of 0 \<Rightarrow> Some x | Suc j \<Rightarrow> xs \<langle>j\<rangle>)"
       
   163 
       
   164 inductive nth_el' :: "'a list \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> bool"
       
   165 where
       
   166   "nth_el' (x # xs) 0 x"
       
   167 | "nth_el' xs i y \<Longrightarrow> nth_el' (x # xs) (Suc i) y"
       
   168 
       
   169 inductive typing :: "type list \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool"  ("_ \<turnstile> _ : _" [50, 50, 50] 50)
       
   170   where
       
   171     Var [intro!]: "nth_el' env x T \<Longrightarrow> env \<turnstile> Var x : T"
       
   172   | Abs [intro!]: "T # env \<turnstile> t : U \<Longrightarrow> env \<turnstile> Abs T t : (T \<Rightarrow> U)"
       
   173   | App [intro!]: "env \<turnstile> s : T \<Rightarrow> U \<Longrightarrow> env \<turnstile> t : T \<Longrightarrow> env \<turnstile> (s \<degree> t) : U"
       
   174 
       
   175 primrec
       
   176   lift :: "[dB, nat] => dB"
       
   177 where
       
   178     "lift (Var i) k = (if i < k then Var i else Var (i + 1))"
       
   179   | "lift (s \<degree> t) k = lift s k \<degree> lift t k"
       
   180   | "lift (Abs T s) k = Abs T (lift s (k + 1))"
       
   181 
       
   182 primrec
       
   183   subst :: "[dB, dB, nat] => dB"  ("_[_'/_]" [300, 0, 0] 300)
       
   184 where
       
   185     subst_Var: "(Var i)[s/k] =
       
   186       (if k < i then Var (i - 1) else if i = k then s else Var i)"
       
   187   | subst_App: "(t \<degree> u)[s/k] = t[s/k] \<degree> u[s/k]"
       
   188   | subst_Abs: "(Abs T t)[s/k] = Abs T (t[lift s 0 / k+1])"
       
   189 
       
   190 inductive beta :: "[dB, dB] => bool"  (infixl "\<rightarrow>\<^sub>\<beta>" 50)
       
   191   where
       
   192     beta [simp, intro!]: "Abs T s \<degree> t \<rightarrow>\<^sub>\<beta> s[t/0]"
       
   193   | appL [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> s \<degree> u \<rightarrow>\<^sub>\<beta> t \<degree> u"
       
   194   | appR [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> u \<degree> s \<rightarrow>\<^sub>\<beta> u \<degree> t"
       
   195   | abs [simp, intro!]: "s \<rightarrow>\<^sub>\<beta> t ==> Abs T s \<rightarrow>\<^sub>\<beta> Abs T t"
       
   196 
       
   197 code_pred (expected_modes: i => i => o => bool, i => i => i => bool) typing .
       
   198 thm typing.equation
       
   199 
       
   200 code_pred (modes: i => i => bool,  i => o => bool as reduce') beta .
       
   201 thm beta.equation
       
   202 
       
   203 values "{x. App (Abs (Atom 0) (Var 0)) (Var 1) \<rightarrow>\<^sub>\<beta> x}"
       
   204 
       
   205 definition "reduce t = Predicate.the (reduce' t)"
       
   206 
       
   207 value "reduce (App (Abs (Atom 0) (Var 0)) (Var 1))"
       
   208 
       
   209 code_pred [dseq] typing .
       
   210 code_pred [random_dseq] typing .
       
   211 
       
   212 values [random_dseq 1,1,5] 10 "{(\<Gamma>, t, T). \<Gamma> \<turnstile> t : T}"
       
   213 
       
   214 subsection {* A minimal example of yet another semantics *}
       
   215 
       
   216 text {* thanks to Elke Salecker *}
       
   217 
       
   218 types
       
   219   vname = nat
       
   220   vvalue = int
       
   221   var_assign = "vname \<Rightarrow> vvalue"  --"variable assignment"
       
   222 
       
   223 datatype ir_expr = 
       
   224   IrConst vvalue
       
   225 | ObjAddr vname
       
   226 | Add ir_expr ir_expr
       
   227 
       
   228 datatype val =
       
   229   IntVal  vvalue
       
   230 
       
   231 record  configuration =
       
   232   Env :: var_assign
       
   233 
       
   234 inductive eval_var ::
       
   235   "ir_expr \<Rightarrow> configuration \<Rightarrow> val \<Rightarrow> bool"
       
   236 where
       
   237   irconst: "eval_var (IrConst i) conf (IntVal i)"
       
   238 | objaddr: "\<lbrakk> Env conf n = i \<rbrakk> \<Longrightarrow> eval_var (ObjAddr n) conf (IntVal i)"
       
   239 | plus: "\<lbrakk> eval_var l conf (IntVal vl); eval_var r conf (IntVal vr) \<rbrakk> \<Longrightarrow> eval_var (Add l r) conf (IntVal (vl+vr))"
       
   240 
       
   241 
       
   242 code_pred eval_var .
       
   243 thm eval_var.equation
       
   244 
       
   245 values "{val. eval_var (Add (IrConst 1) (IrConst 2)) (| Env = (\<lambda>x. 0)|) val}"
       
   246 
       
   247 subsection {* Another semantics *}
       
   248 
       
   249 types
       
   250   name = nat --"For simplicity in examples"
       
   251   state' = "name \<Rightarrow> nat"
       
   252 
       
   253 datatype aexp = N nat | V name | Plus aexp aexp
       
   254 
       
   255 fun aval :: "aexp \<Rightarrow> state' \<Rightarrow> nat" where
       
   256 "aval (N n) _ = n" |
       
   257 "aval (V x) st = st x" |
       
   258 "aval (Plus e\<^isub>1 e\<^isub>2) st = aval e\<^isub>1 st + aval e\<^isub>2 st"
       
   259 
       
   260 datatype bexp = B bool | Not bexp | And bexp bexp | Less aexp aexp
       
   261 
       
   262 primrec bval :: "bexp \<Rightarrow> state' \<Rightarrow> bool" where
       
   263 "bval (B b) _ = b" |
       
   264 "bval (Not b) st = (\<not> bval b st)" |
       
   265 "bval (And b1 b2) st = (bval b1 st \<and> bval b2 st)" |
       
   266 "bval (Less a\<^isub>1 a\<^isub>2) st = (aval a\<^isub>1 st < aval a\<^isub>2 st)"
       
   267 
       
   268 datatype
       
   269   com' = SKIP 
       
   270       | Assign name aexp         ("_ ::= _" [1000, 61] 61)
       
   271       | Semi   com'  com'          ("_; _"  [60, 61] 60)
       
   272       | If     bexp com' com'     ("IF _ THEN _ ELSE _"  [0, 0, 61] 61)
       
   273       | While  bexp com'         ("WHILE _ DO _"  [0, 61] 61)
       
   274 
       
   275 inductive
       
   276   big_step :: "com' * state' \<Rightarrow> state' \<Rightarrow> bool" (infix "\<Rightarrow>" 55)
       
   277 where
       
   278   Skip:    "(SKIP,s) \<Rightarrow> s"
       
   279 | Assign:  "(x ::= a,s) \<Rightarrow> s(x := aval a s)"
       
   280 
       
   281 | Semi:    "(c\<^isub>1,s\<^isub>1) \<Rightarrow> s\<^isub>2  \<Longrightarrow>  (c\<^isub>2,s\<^isub>2) \<Rightarrow> s\<^isub>3  \<Longrightarrow> (c\<^isub>1;c\<^isub>2, s\<^isub>1) \<Rightarrow> s\<^isub>3"
       
   282 
       
   283 | IfTrue:  "bval b s  \<Longrightarrow>  (c\<^isub>1,s) \<Rightarrow> t  \<Longrightarrow>  (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t"
       
   284 | IfFalse: "\<not>bval b s  \<Longrightarrow>  (c\<^isub>2,s) \<Rightarrow> t  \<Longrightarrow>  (IF b THEN c\<^isub>1 ELSE c\<^isub>2, s) \<Rightarrow> t"
       
   285 
       
   286 | WhileFalse: "\<not>bval b s \<Longrightarrow> (WHILE b DO c,s) \<Rightarrow> s"
       
   287 | WhileTrue:  "bval b s\<^isub>1  \<Longrightarrow>  (c,s\<^isub>1) \<Rightarrow> s\<^isub>2  \<Longrightarrow>  (WHILE b DO c, s\<^isub>2) \<Rightarrow> s\<^isub>3
       
   288                \<Longrightarrow> (WHILE b DO c, s\<^isub>1) \<Rightarrow> s\<^isub>3"
       
   289 
       
   290 code_pred big_step .
       
   291 
       
   292 thm big_step.equation
       
   293 
       
   294 subsection {* CCS *}
       
   295 
       
   296 text{* This example formalizes finite CCS processes without communication or
       
   297 recursion. For simplicity, labels are natural numbers. *}
       
   298 
       
   299 datatype proc = nil | pre nat proc | or proc proc | par proc proc
       
   300 
       
   301 inductive step :: "proc \<Rightarrow> nat \<Rightarrow> proc \<Rightarrow> bool" where
       
   302 "step (pre n p) n p" |
       
   303 "step p1 a q \<Longrightarrow> step (or p1 p2) a q" |
       
   304 "step p2 a q \<Longrightarrow> step (or p1 p2) a q" |
       
   305 "step p1 a q \<Longrightarrow> step (par p1 p2) a (par q p2)" |
       
   306 "step p2 a q \<Longrightarrow> step (par p1 p2) a (par p1 q)"
       
   307 
       
   308 code_pred step .
       
   309 
       
   310 inductive steps where
       
   311 "steps p [] p" |
       
   312 "step p a q \<Longrightarrow> steps q as r \<Longrightarrow> steps p (a#as) r"
       
   313 
       
   314 code_pred steps .
       
   315 
       
   316 values 3 
       
   317  "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
       
   318 
       
   319 values 5
       
   320  "{as . steps (par (or (pre 0 nil) (pre 1 nil)) (pre 2 nil)) as (par nil nil)}"
       
   321 
       
   322 values 3 "{(a,q). step (par nil nil) a q}"
       
   323 
       
   324 
       
   325 end
       
   326