# HG changeset patch # User lcp # Date 793899998 -3600 # Node ID b6c0407f203e716e48cf78a9481b76ffdf420c3e # Parent 12943ab62cc55d655e582e62d6f0ede83c9bbf13 Installation of Integ (ported from ZF by Mattolini) diff -r 12943ab62cc5 -r b6c0407f203e Integ/Integ.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Integ/Integ.thy Mon Feb 27 16:46:38 1995 +0100 @@ -0,0 +1,80 @@ +(* Title: Integ.thy + ID: $Id$ + Authors: Riccardo Mattolini, Dip. Sistemi e Informatica + Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 Universita' di Firenze + Copyright 1993 University of Cambridge + +The integers as equivalence classes over nat*nat. +*) + +Integ = Equiv + Arith + +consts + intrel :: "((nat * nat) * (nat * nat)) set" + +defs + intrel_def + "intrel == {p. ? x1 y1 x2 y2. p=<,> & x1+y2 = x2+y1}" + +subtype (Integ) + int = "{x::(nat*nat).True}/intrel" ("quotient_def") + +arities int :: ord + int :: plus + int :: times + int :: minus +consts + zNat :: "nat set" + znat :: "nat => int" ("$# _" [80] 80) + zminus :: "int => int" ("$~ _" [80] 80) + znegative :: "int => bool" + zmagnitude :: "int => int" + zdiv,zmod :: "[int,int]=>int" (infixl 70) + zpred,zsuc :: "int=>int" + +defs + zNat_def "zNat == {x::nat. True}" + + znat_def "$# m == Abs_Integ(intrel ^^ {})" + + zminus_def + "$~ Z == Abs_Integ(UN p:Rep_Integ(Z). split(%x y. intrel^^{},p))" + + znegative_def + "znegative(Z) == EX x y. x:Rep_Integ(Z)" + + zmagnitude_def + "zmagnitude(Z) == Abs_Integ(UN p:Rep_Integ(Z).split(%x y. intrel^^{<(y-x) + (x-y),0>},p))" + + zadd_def + "Z1 + Z2 == \ +\ Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). \ +\ split(%x1 y1. split(%x2 y2. intrel^^{},p2),p1))" + + zdiff_def "Z1 - Z2 == Z1 + zminus(Z2)" + + zless_def "Z1},p2),p1))" + + zdiv_def + "Z1 zdiv Z2 == \ +\ Abs_Integ(UN p1:Rep_Integ(Z1). UN p2:Rep_Integ(Z2). split(%x1 y1. \ +\ split(%x2 y2. intrel^^{<(x1-y1)div(x2-y2)+(y1-x1)div(y2-x2), \ +\ (x1-y1)div(y2-x2)+(y1-x1)div(x2-y2)>},p2),p1))" + + zmod_def + "Z1 zmod Z2 == \ +\ Abs_Integ(UN p1:Rep_Integ(Z1).UN p2:Rep_Integ(Z2).split(%x1 y1. \ +\ split(%x2 y2. intrel^^{<(x1-y1)mod((x2-y2)+(y2-x2)), \ +\ (x1-y1)mod((x2-y2)+(x2-y2))>},p2),p1))" + + zsuc_def "zsuc(Z) == Z + $# Suc(0)" + + zpred_def "zpred(Z) == Z - $# Suc(0)" +end diff -r 12943ab62cc5 -r b6c0407f203e Integ/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Integ/ROOT.ML Mon Feb 27 16:46:38 1995 +0100 @@ -0,0 +1,12 @@ +(* Title: HOL/Integ/ROOT + ID: $Id$ + Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1995 University of Cambridge + +The Integers in HOL (ported from ZF by Riccardo Mattolini) +*) + +HOL_build_completed; (*Cause examples to fail if HOL did*) + +loadpath := ["Integ"]; +time_use_thy "Integ" handle _ => exit 1; diff -r 12943ab62cc5 -r b6c0407f203e Integ/Relation.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Integ/Relation.ML Mon Feb 27 16:46:38 1995 +0100 @@ -0,0 +1,98 @@ +(* Title: Relation.ML + ID: $Id$ + Authors: Riccardo Mattolini, Dip. Sistemi e Informatica + Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 Universita' di Firenze + Copyright 1993 University of Cambridge + +Functions represented as relations in HOL Set Theory +*) + +val RSLIST = curry (op MRS); + +open Relation; + +goalw Relation.thy [converse_def] "!!a b r. :r ==> :converse(r)"; +by (simp_tac prod_ss 1); +by (fast_tac set_cs 1); +qed "converseI"; + +goalw Relation.thy [converse_def] "!!a b r. : converse(r) ==> : r"; +by (fast_tac comp_cs 1); +qed "converseD"; + +qed_goalw "converseE" Relation.thy [converse_def] + "[| yx : converse(r); \ +\ !!x y. [| yx=; :r |] ==> P \ +\ |] ==> P" + (fn [major,minor]=> + [ (rtac (major RS CollectE) 1), + (REPEAT (eresolve_tac [bexE,exE, conjE, minor] 1)), + (hyp_subst_tac 1), + (assume_tac 1) ]); + +val converse_cs = comp_cs addSIs [converseI] + addSEs [converseD,converseE]; + +qed_goalw "Domain_iff" Relation.thy [Domain_def] + "a: Domain(r) = (EX y. : r)" + (fn _=> [ (fast_tac comp_cs 1) ]); + +qed_goal "DomainI" Relation.thy "!!a b r. : r ==> a: Domain(r)" + (fn _ => [ (etac (exI RS (Domain_iff RS iffD2)) 1) ]); + +qed_goal "DomainE" Relation.thy + "[| a : Domain(r); !!y. : r ==> P |] ==> P" + (fn prems=> + [ (rtac (Domain_iff RS iffD1 RS exE) 1), + (REPEAT (ares_tac prems 1)) ]); + +qed_goalw "RangeI" Relation.thy [Range_def] "!!a b r.: r ==> b : Range(r)" + (fn _ => [ (etac (converseI RS DomainI) 1) ]); + +qed_goalw "RangeE" Relation.thy [Range_def] + "[| b : Range(r); !!x. : r ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS DomainE) 1), + (resolve_tac prems 1), + (etac converseD 1) ]); + +(*** Image of a set under a function/relation ***) + +qed_goalw "Image_iff" Relation.thy [Image_def] + "b : r^^A = (? x:A. :r)" + (fn _ => [ fast_tac (comp_cs addIs [RangeI]) 1 ]); + +qed_goal "Image_singleton_iff" Relation.thy + "(b : r^^{a}) = (:r)" + (fn _ => [ rtac (Image_iff RS trans) 1, + fast_tac comp_cs 1 ]); + +qed_goalw "ImageI" Relation.thy [Image_def] + "!!a b r. [| : r; a:A |] ==> b : r^^A" + (fn _ => [ (REPEAT (ares_tac [CollectI,RangeI,bexI] 1)), + (resolve_tac [conjI ] 1), + (resolve_tac [RangeI] 1), + (REPEAT (fast_tac set_cs 1))]); + +qed_goalw "ImageE" Relation.thy [Image_def] + "[| b: r^^A; !!x.[| : r; x:A |] ==> P |] ==> P" + (fn major::prems=> + [ (rtac (major RS CollectE) 1), + (safe_tac set_cs), + (etac RangeE 1), + (rtac (hd prems) 1), + (REPEAT (etac bexE 1 ORELSE ares_tac prems 1)) ]); + +qed_goal "Image_subset" Relation.thy + "!!A B r. r <= Sigma(A,%x.B) ==> r^^C <= B" + (fn _ => + [ (rtac subsetI 1), + (REPEAT (eresolve_tac [asm_rl, ImageE, subsetD RS SigmaD2] 1)) ]); + +val rel_cs = converse_cs addSIs [converseI] + addIs [ImageI, DomainI, RangeI] + addSEs [ImageE, DomainE, RangeE]; + +val rel_eq_cs = rel_cs addSIs [equalityI]; + diff -r 12943ab62cc5 -r b6c0407f203e Integ/Relation.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Integ/Relation.thy Mon Feb 27 16:46:38 1995 +0100 @@ -0,0 +1,24 @@ +(* Title: Relation.thy + ID: $Id$ + Author: Riccardo Mattolini, Dip. Sistemi e Informatica + and Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1994 Universita' di Firenze + Copyright 1993 University of Cambridge + +Functions represented as relations in Higher-Order Set Theory +*) + +Relation = Trancl + +consts + converse :: "('a*'a) set => ('a*'a) set" + "^^" :: "[('a*'a) set,'a set] => 'a set" (infixl 90) + Domain :: "('a*'a) set => 'a set" + Range :: "('a*'a) set => 'a set" + +defs + converse_def "converse(r) == {z. (? w:r. ? x y. w= & z=)}" + Domain_def "Domain(r) == {z. ! x. (z=x --> (? y. :r))}" + Range_def "Range(r) == Domain(converse(r))" + Image_def "r ^^ s == {y. y:Range(r) & (? x:s. :r)}" + +end