author | nipkow |
Mon, 06 Dec 1999 14:23:45 +0100 | |
changeset 8048 | b7f7e18eb584 |
parent 8011 | d14c4e9e9c8e |
child 8116 | 759f712f8f06 |
permissions | -rw-r--r-- |
(* Title: HOL/MicroJava/J/JBasis.thy ID: $Id$ Author: David von Oheimb Copyright 1999 TU Muenchen Some auxiliary definitions. *) JBasis = Main + constdefs unique :: "('a \\<times> 'b) list \\<Rightarrow> bool" "unique \\<equiv> nodups \\<circ> map fst" list_all2 :: "('a \\<Rightarrow> 'b \\<Rightarrow> bool) \\<Rightarrow> ('a list \\<Rightarrow> 'b list \\<Rightarrow> bool)" "list_all2 P xs ys \\<equiv> length xs = length ys \\<and> (\\<forall> (x,y)\\<in>set (zip xs ys). P x y)" end