| author | wenzelm | 
| Wed, 24 May 2017 11:17:23 +0200 | |
| changeset 65915 | 49f61e2f5a02 | 
| parent 64593 | 50c715579715 | 
| child 66453 | cc19f7ca2ed6 | 
| permissions | -rw-r--r-- | 
| 45044 | 1 | theory Proc1 | 
| 58130 | 2 | imports "../SPARK" | 
| 45044 | 3 | begin | 
| 4 | ||
| 56798 
939e88e79724
Discontinued old spark_open; spark_open_siv is now spark_open
 berghofe parents: 
45044diff
changeset | 5 | spark_open "loop_invariant/proc1" | 
| 45044 | 6 | |
| 7 | spark_vc procedure_proc1_5 | |
| 64593 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
58130diff
changeset | 8 | by (simp add: ring_distribs mod_simps) | 
| 45044 | 9 | |
| 10 | spark_vc procedure_proc1_8 | |
| 64593 
50c715579715
reoriented congruence rules in non-explosive direction
 haftmann parents: 
58130diff
changeset | 11 | by (simp add: ring_distribs mod_simps) | 
| 45044 | 12 | |
| 13 | spark_end | |
| 14 | ||
| 15 | lemma pow_2_32_simp: "4294967296 = (2::int)^32" | |
| 16 | by simp | |
| 17 | ||
| 18 | end |