src/HOL/SPARK/Manual/Proc1.thy
author wenzelm
Fri, 18 Aug 2017 20:47:47 +0200
changeset 66453 cc19f7ca2ed6
parent 64593 50c715579715
child 66992 69673025292e
permissions -rw-r--r--
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45044
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
     1
theory Proc1
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 64593
diff changeset
     2
imports SPARK
45044
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
     3
begin
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
     4
56798
939e88e79724 Discontinued old spark_open; spark_open_siv is now spark_open
berghofe
parents: 45044
diff changeset
     5
spark_open "loop_invariant/proc1"
45044
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
     6
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
     7
spark_vc procedure_proc1_5
64593
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 58130
diff changeset
     8
  by (simp add: ring_distribs mod_simps)
45044
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
     9
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
    10
spark_vc procedure_proc1_8
64593
50c715579715 reoriented congruence rules in non-explosive direction
haftmann
parents: 58130
diff changeset
    11
  by (simp add: ring_distribs mod_simps)
45044
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
    12
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
    13
spark_end
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
    14
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
    15
lemma pow_2_32_simp: "4294967296 = (2::int)^32"
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
    16
  by simp
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
    17
2fae15f8984d Added documentation for HOL-SPARK
berghofe
parents:
diff changeset
    18
end