src/HOLCF/Cprod1.ML
author berghofe
Fri, 11 Jul 2003 14:55:17 +0200
changeset 14102 8af7334af4b3
parent 12030 46d57d0290a2
child 14981 e73f8140af78
permissions -rw-r--r--
- Installed specific code generator for equality enforcing that arguments do not have function types, which would result in an error message during compilation. - Added test case generators for basic types.

(*  Title:      HOLCF/Cprod1.ML
    ID:         $Id$
    Author:     Franz Regensburger
    License:    GPL (GNU GENERAL PUBLIC LICENSE)

Partial ordering for cartesian product of HOL theory Product_Type.thy
*)


(* ------------------------------------------------------------------------ *)
(* less_cprod is a partial order on 'a * 'b                                 *)
(* ------------------------------------------------------------------------ *)

Goalw [less_cprod_def] "(p::'a*'b) << p";
by (Simp_tac 1);
qed "refl_less_cprod";

Goalw [less_cprod_def] "[|(p1::'a * 'b) << p2;p2 << p1|] ==> p1=p2";
by (rtac injective_fst_snd 1);
by (fast_tac (HOL_cs addIs [antisym_less]) 1);
by (fast_tac (HOL_cs addIs [antisym_less]) 1);
qed "antisym_less_cprod";

Goalw [less_cprod_def]
        "[|(p1::'a*'b) << p2;p2 << p3|] ==> p1 << p3";
by (rtac conjI 1);
by (fast_tac (HOL_cs addIs [trans_less]) 1);
by (fast_tac (HOL_cs addIs [trans_less]) 1);
qed "trans_less_cprod";