# HG changeset patch # User nipkow # Date 1035811837 -3600 # Node ID 47fdf4e8e89c600046817f97d6cd6bc991b53b62 # Parent 91674c8a008b5ee1ed949cb80fae1280c00a9dca *** empty log message *** diff -r 91674c8a008b -r 47fdf4e8e89c src/HOL/Hoare/Examples.ML --- a/src/HOL/Hoare/Examples.ML Mon Oct 28 14:29:51 2002 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,207 +0,0 @@ -(* Title: HOL/Hoare/Examples.ML - ID: $Id$ - Author: Norbert Galm & Tobias Nipkow - Copyright 1998 TUM -*) - -(*** ARITHMETIC ***) - -(** multiplication by successive addition **) - -Goal "|- VARS m s a b. \ -\ {a=A & b=B} \ -\ m := 0; s := 0; \ -\ WHILE m~=a \ -\ INV {s=m*b & a=A & b=B} \ -\ DO s := s+b; m := m+(1::nat) OD \ -\ {s = A*B}"; -by (hoare_tac (Asm_full_simp_tac) 1); -qed "multiply_by_add"; - -(** Euclid's algorithm for GCD **) - -Goal "|- VARS a b. \ -\ {0 A!j ~= key} \ -\ DO i := i+1 OD \ -\ {(i < length A --> A!i = key) & \ -\ (i = length A --> (!j. j < length A --> A!j ~= key))}"; -by (hoare_tac Asm_full_simp_tac 1); -by (blast_tac (claset() addSEs [less_SucE]) 1); -qed "zero_search"; - -(* -The `partition' procedure for quicksort. -`A' is the array to be sorted (modelled as a list). -Elements of A must be of class order to infer at the end -that the elements between u and l are equal to pivot. - -Ambiguity warnings of parser are due to := being used -both for assignment and list update. -*) -Goal "m - Suc 0 < n ==> m < Suc n"; -by (arith_tac 1); -qed "lemma"; - -Goal -"[| leq == %A i. !k. k A!k <= pivot; \ -\ geq == %A i. !k. i pivot <= A!k |] ==> \ -\ |- VARS A u l.\ -\ {0 < length(A::('a::order)list)} \ -\ l := 0; u := length A - Suc 0; \ -\ WHILE l <= u \ -\ INV {leq A l & geq A u & u A!k = pivot) & geq A l}"; -(* expand and delete abbreviations first *) -by (asm_simp_tac HOL_basic_ss 1); -by (REPEAT (etac thin_rl 1)); -by (hoare_tac Asm_full_simp_tac 1); - by (force_tac (claset(), simpset() addsimps [neq_Nil_conv]) 1); - by (blast_tac (claset() addSEs [less_SucE] addIs [Suc_leI]) 1); - by (blast_tac (claset() addSEs [less_SucE] - addIs [less_imp_diff_less] addDs [lemma]) 1); - by (force_tac (claset(), simpset() addsimps [nth_list_update]) 1); -by (auto_tac (claset() addIs [order_antisym], simpset())); -qed "Partition";