src/HOL/ex/Points.ML
author wenzelm
Fri, 23 Oct 1998 20:28:33 +0200
changeset 5753 c90b5f7d0c61
parent 5741 139a25a1e01e
child 5846 d99feda2d226
permissions -rw-r--r--
tuned;


(*Note: any of these goals may be solved by a single stroke of
  auto(); or force 1;*)


(* some basic simplifications *)

Goal "!!m n p. x (| x = m, y = n, ... = p |) = m";
by (Simp_tac 1);
result();

Goal "!!m n p. (| x = m, y = n, ... = p |) (| x:= 0 |) = (| x = 0, y = n, ... = p |)";
by (Simp_tac 1);
result();


(* splitting quantifiers *)

Goal "!!r m n. r (| x := n |) (| y := m |) = r (| y := m |) (| x := n |)";
by (record_split_tac 1);
by (Simp_tac 1);
result();

Goal "!!r m n. r (| x := n |) (| x := m |) = r (| x := m |)";
by (record_split_tac 1);
by (Simp_tac 1);
result();


(* record equality *)

Goal "!!r. (| x = n, y = p |) = (| x = n', y = p' |) & n = 0 ==> n' = 0";
by (Asm_full_simp_tac 1);
by (Blast_tac 1);
result();

Goalw [getX_def, setX_def, incX_def] "!!r n. incX r = setX r (Suc (getX r))";
by (record_split_tac 1);
by (Simp_tac 1);
result();