# HG changeset patch # User paulson # Date 1042645461 -3600 # Node ID 2a34dc5cf79e0a9998e99b0c307cc260ead231f4 # Parent 61272514e3b5a34d50a2801bd9c2de37f57a4dde moving "let" from ZF to FOL diff -r 61272514e3b5 -r 2a34dc5cf79e src/FOL/IFOL.thy --- a/src/FOL/IFOL.thy Wed Jan 15 16:43:12 2003 +0100 +++ b/src/FOL/IFOL.thy Wed Jan 15 16:44:21 2003 +0100 @@ -112,6 +112,7 @@ iff_reflection: "(P<->Q) ==> (P==Q)" + subsection {* Lemmas and proof tools *} setup Simplifier.setup @@ -244,4 +245,38 @@ mp trans + + +subsection {* ``Let'' declarations *} + +nonterminals letbinds letbind + +constdefs + Let :: "['a::logic, 'a => 'b] => ('b::logic)" + "Let(s, f) == f(s)" + +syntax + "_bind" :: "[pttrn, 'a] => letbind" ("(2_ =/ _)" 10) + "" :: "letbind => letbinds" ("_") + "_binds" :: "[letbind, letbinds] => letbinds" ("_;/ _") + "_Let" :: "[letbinds, 'a] => 'a" ("(let (_)/ in (_))" 10) + +translations + "_Let(_binds(b, bs), e)" == "_Let(b, _Let(bs, e))" + "let x = a in e" == "Let(a, %x. e)" + + +lemma LetI: + assumes prem: "(!!x. x=t ==> P(u(x)))" + shows "P(let x=t in u(x))" +apply (unfold Let_def) +apply (rule refl [THEN prem]) +done + +ML +{* +val Let_def = thm "Let_def"; +val LetI = thm "LetI"; +*} + end