src/HOL/Predicate_Compile_Examples/Examples.thy
changeset 42094 e6867e9c6d10
parent 41413 64cd30d6b0b8
child 42187 b4f4ed5b8586
equal deleted inserted replaced
42093:85f487b8e70c 42094:e6867e9c6d10
   289 
   289 
   290 code_pred big_step .
   290 code_pred big_step .
   291 
   291 
   292 thm big_step.equation
   292 thm big_step.equation
   293 
   293 
       
   294 definition list :: "(nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a list" where
       
   295   "list s n = map s [0 ..< n]"
       
   296 
       
   297 values [expected "{[42, (43 :: nat)]}"] "{list s 2|s. (SKIP, nth [42, 43]) \<Rightarrow> s}"
       
   298 
       
   299 
   294 subsection {* CCS *}
   300 subsection {* CCS *}
   295 
   301 
   296 text{* This example formalizes finite CCS processes without communication or
   302 text{* This example formalizes finite CCS processes without communication or
   297 recursion. For simplicity, labels are natural numbers. *}
   303 recursion. For simplicity, labels are natural numbers. *}
   298 
   304