# HG changeset patch # User haftmann # Date 1267447351 -3600 # Node ID 47ee18b6ae324891b95a24206e6612500f52e08c # Parent 1810b1ade437883b807ab6cf3504dbe8a2153a9d# Parent d8d7d1b785afa9a54b69c7326b5f11c90401e804 merged diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/TutorialI/Advanced/Partial.thy --- a/doc-src/TutorialI/Advanced/Partial.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/TutorialI/Advanced/Partial.thy Mon Mar 01 13:42:31 2010 +0100 @@ -34,7 +34,7 @@ preconditions: *} -constdefs subtract :: "nat \ nat \ nat" +definition subtract :: "nat \ nat \ nat" where "n \ m \ subtract m n \ m - n" text{* @@ -85,7 +85,7 @@ Phrased differently, the relation *} -constdefs step1 :: "('a \ 'a) \ ('a \ 'a)set" +definition step1 :: "('a \ 'a) \ ('a \ 'a)set" where "step1 f \ {(y,x). y = f x \ y \ x}" text{*\noindent @@ -160,7 +160,7 @@ consider the following definition of function @{const find}: *} -constdefs find2 :: "('a \ 'a) \ 'a \ 'a" +definition find2 :: "('a \ 'a) \ 'a \ 'a" where "find2 f x \ fst(while (\(x,x'). x' \ x) (\(x,x'). (x',f x')) (x,f x))" diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/TutorialI/CTL/CTL.thy --- a/doc-src/TutorialI/CTL/CTL.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/TutorialI/CTL/CTL.thy Mon Mar 01 13:42:31 2010 +0100 @@ -365,8 +365,7 @@ *} (*<*) -constdefs - eufix :: "state set \ state set \ state set \ state set" +definition eufix :: "state set \ state set \ state set \ state set" where "eufix A B T \ B \ A \ (M\ `` T)" lemma "lfp(eufix A B) \ eusem A B" @@ -397,8 +396,7 @@ done (* -constdefs - eusem :: "state set \ state set \ state set" +definition eusem :: "state set \ state set \ state set" where "eusem A B \ {s. \p\Paths s. \j. p j \ B \ (\i < j. p i \ A)}" axioms @@ -414,8 +412,7 @@ apply(blast intro: M_total[THEN someI_ex]) done -constdefs - pcons :: "state \ (nat \ state) \ (nat \ state)" +definition pcons :: "state \ (nat \ state) \ (nat \ state)" where "pcons s p == \i. case i of 0 \ s | Suc j \ p j" lemma pcons_PathI: "[| (s,t) : M; p \ Paths t |] ==> pcons s p \ Paths s"; diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/TutorialI/Misc/Option2.thy --- a/doc-src/TutorialI/Misc/Option2.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/TutorialI/Misc/Option2.thy Mon Mar 01 13:42:31 2010 +0100 @@ -24,8 +24,7 @@ *} (*<*) (* -constdefs - infplus :: "nat option \ nat option \ nat option" +definition infplus :: "nat option \ nat option \ nat option" where "infplus x y \ case x of None \ None | Some m \ (case y of None \ None | Some n \ Some(m+n))" diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/TutorialI/Overview/LNCS/FP1.thy --- a/doc-src/TutorialI/Overview/LNCS/FP1.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/TutorialI/Overview/LNCS/FP1.thy Mon Mar 01 13:42:31 2010 +0100 @@ -62,7 +62,7 @@ consts xor :: "bool \ bool \ bool" defs xor_def: "xor x y \ x \ \y \ \x \ y" -constdefs nand :: "bool \ bool \ bool" +definition nand :: "bool \ bool \ bool" where "nand x y \ \(x \ y)" lemma "\ xor x x" diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/TutorialI/Overview/LNCS/Ordinal.thy --- a/doc-src/TutorialI/Overview/LNCS/Ordinal.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/TutorialI/Overview/LNCS/Ordinal.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,8 +9,7 @@ "pred (Succ a) n = Some a" "pred (Limit f) n = Some (f n)" -constdefs - OpLim :: "(nat \ (ordinal \ ordinal)) \ (ordinal \ ordinal)" +definition OpLim :: "(nat \ (ordinal \ ordinal)) \ (ordinal \ ordinal)" where "OpLim F a \ Limit (\n. F n a)" OpItw :: "(ordinal \ ordinal) \ (ordinal \ ordinal)" ("\") "\f \ OpLim (power f)" @@ -29,8 +28,7 @@ "\f (Succ a) = f (Succ (\f a))" "\f (Limit h) = Limit (\n. \f (h n))" -constdefs - deriv :: "(ordinal \ ordinal) \ (ordinal \ ordinal)" +definition deriv :: "(ordinal \ ordinal) \ (ordinal \ ordinal)" where "deriv f \ \(\f)" consts @@ -40,8 +38,7 @@ "veblen (Succ a) = \(OpLim (power (veblen a)))" "veblen (Limit f) = \(OpLim (\n. veblen (f n)))" -constdefs - veb :: "ordinal \ ordinal" +definition veb :: "ordinal \ ordinal" where "veb a \ veblen a Zero" epsilon0 :: ordinal ("\\<^sub>0") "\\<^sub>0 \ veb Zero" diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/TutorialI/Protocol/Message.thy --- a/doc-src/TutorialI/Protocol/Message.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/TutorialI/Protocol/Message.thy Mon Mar 01 13:42:31 2010 +0100 @@ -46,8 +46,7 @@ text{*The inverse of a symmetric key is itself; that of a public key is the private key and vice versa*} -constdefs - symKeys :: "key set" +definition symKeys :: "key set" where "symKeys == {K. invKey K = K}" (*>*) @@ -92,8 +91,7 @@ "{|x, y|}" == "CONST MPair x y" -constdefs - keysFor :: "msg set => key set" +definition keysFor :: "msg set => key set" where --{*Keys useful to decrypt elements of a message set*} "keysFor H == invKey ` {K. \X. Crypt K X \ H}" diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/TutorialI/Rules/Primes.thy --- a/doc-src/TutorialI/Rules/Primes.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/TutorialI/Rules/Primes.thy Mon Mar 01 13:42:31 2010 +0100 @@ -99,8 +99,7 @@ (**** The material below was omitted from the book ****) -constdefs - is_gcd :: "[nat,nat,nat] \ bool" (*gcd as a relation*) +definition is_gcd :: "[nat,nat,nat] \ bool" where (*gcd as a relation*) "is_gcd p m n == p dvd m \ p dvd n \ (ALL d. d dvd m \ d dvd n \ d dvd p)" diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/TutorialI/Sets/Examples.thy --- a/doc-src/TutorialI/Sets/Examples.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/TutorialI/Sets/Examples.thy Mon Mar 01 13:42:31 2010 +0100 @@ -156,8 +156,7 @@ lemma "{x. P x \ Q x} = -{x. P x} \ {x. Q x}" by blast -constdefs - prime :: "nat set" +definition prime :: "nat set" where "prime == {p. 1

m=1 | m=p)}" lemma "{p*q | p q. p\prime \ q\prime} = diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/ZF/If.thy --- a/doc-src/ZF/If.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/ZF/If.thy Mon Mar 01 13:42:31 2010 +0100 @@ -8,8 +8,7 @@ theory If imports FOL begin -constdefs - "if" :: "[o,o,o]=>o" +definition "if" :: "[o,o,o]=>o" where "if(P,Q,R) == P&Q | ~P&R" lemma ifI: diff -r 1810b1ade437 -r 47ee18b6ae32 doc-src/ZF/ZF_examples.thy --- a/doc-src/ZF/ZF_examples.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/doc-src/ZF/ZF_examples.thy Mon Mar 01 13:42:31 2010 +0100 @@ -64,7 +64,7 @@ "t \ bt(A) ==> \k \ nat. n_nodes_aux(t)`k = n_nodes(t) #+ k" by (induct_tac t, simp_all) -constdefs n_nodes_tail :: "i => i" +definition n_nodes_tail :: "i => i" where "n_nodes_tail(t) == n_nodes_aux(t) ` 0" lemma "t \ bt(A) ==> n_nodes_tail(t) = n_nodes(t)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/FOL/IFOL.thy --- a/src/FOL/IFOL.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/FOL/IFOL.thy Mon Mar 01 13:42:31 2010 +0100 @@ -760,8 +760,7 @@ nonterminals letbinds letbind -constdefs - Let :: "['a::{}, 'a => 'b] => ('b::{})" +definition Let :: "['a::{}, 'a => 'b] => ('b::{})" where "Let(s, f) == f(s)" syntax diff -r 1810b1ade437 -r 47ee18b6ae32 src/FOL/ex/If.thy --- a/src/FOL/ex/If.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/FOL/ex/If.thy Mon Mar 01 13:42:31 2010 +0100 @@ -7,8 +7,7 @@ theory If imports FOL begin -constdefs - "if" :: "[o,o,o]=>o" +definition "if" :: "[o,o,o]=>o" where "if(P,Q,R) == P&Q | ~P&R" lemma ifI: diff -r 1810b1ade437 -r 47ee18b6ae32 src/FOLP/ex/If.thy --- a/src/FOLP/ex/If.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/FOLP/ex/If.thy Mon Mar 01 13:42:31 2010 +0100 @@ -4,8 +4,7 @@ imports FOLP begin -constdefs - "if" :: "[o,o,o]=>o" +definition "if" :: "[o,o,o]=>o" where "if(P,Q,R) == P&Q | ~P&R" lemma ifI: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Algebra/AbelCoset.thy --- a/src/HOL/Algebra/AbelCoset.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Algebra/AbelCoset.thy Mon Mar 01 13:42:31 2010 +0100 @@ -38,15 +38,12 @@ ("racong\ _") "a_r_congruent G \ r_congruent \carrier = carrier G, mult = add G, one = zero G\" -constdefs - A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \ ('a set) monoid" - (infixl "A'_Mod" 65) +definition A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \ ('a set) monoid" (infixl "A'_Mod" 65) where --{*Actually defined for groups rather than monoids*} "A_FactGroup G H \ FactGroup \carrier = carrier G, mult = add G, one = zero G\ H" -constdefs - a_kernel :: "('a, 'm) ring_scheme \ ('b, 'n) ring_scheme \ - ('a \ 'b) \ 'a set" +definition a_kernel :: "('a, 'm) ring_scheme \ ('b, 'n) ring_scheme \ + ('a \ 'b) \ 'a set" where --{*the kernel of a homomorphism (additive)*} "a_kernel G H h \ kernel \carrier = carrier G, mult = add G, one = zero G\ \carrier = carrier H, mult = add H, one = zero H\ h" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Algebra/Bij.thy --- a/src/HOL/Algebra/Bij.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Algebra/Bij.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Algebra/Bij.thy - ID: $Id$ Author: Florian Kammueller, with new proofs by L C Paulson *) @@ -8,12 +7,11 @@ section {* Bijections of a Set, Permutation and Automorphism Groups *} -constdefs - Bij :: "'a set \ ('a \ 'a) set" +definition Bij :: "'a set \ ('a \ 'a) set" where --{*Only extensional functions, since otherwise we get too many.*} "Bij S \ extensional S \ {f. bij_betw f S S}" - BijGroup :: "'a set \ ('a \ 'a) monoid" +definition BijGroup :: "'a set \ ('a \ 'a) monoid" where "BijGroup S \ \carrier = Bij S, mult = \g \ Bij S. \f \ Bij S. compose S g f, @@ -71,11 +69,10 @@ done -constdefs - auto :: "('a, 'b) monoid_scheme \ ('a \ 'a) set" +definition auto :: "('a, 'b) monoid_scheme \ ('a \ 'a) set" where "auto G \ hom G G \ Bij (carrier G)" - AutoGroup :: "('a, 'c) monoid_scheme \ ('a \ 'a) monoid" +definition AutoGroup :: "('a, 'c) monoid_scheme \ ('a \ 'a) monoid" where "AutoGroup G \ BijGroup (carrier G) \carrier := auto G\" lemma (in group) id_in_auto: "(\x \ carrier G. x) \ auto G" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Algebra/Coset.thy --- a/src/HOL/Algebra/Coset.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Algebra/Coset.thy Mon Mar 01 13:42:31 2010 +0100 @@ -751,8 +751,7 @@ subsection {*Order of a Group and Lagrange's Theorem*} -constdefs - order :: "('a, 'b) monoid_scheme \ nat" +definition order :: "('a, 'b) monoid_scheme \ nat" where "order S \ card (carrier S)" lemma (in group) rcosets_part_G: @@ -822,9 +821,7 @@ subsection {*Quotient Groups: Factorization of a Group*} -constdefs - FactGroup :: "[('a,'b) monoid_scheme, 'a set] \ ('a set) monoid" - (infixl "Mod" 65) +definition FactGroup :: "[('a,'b) monoid_scheme, 'a set] \ ('a set) monoid" (infixl "Mod" 65) where --{*Actually defined for groups rather than monoids*} "FactGroup G H \ \carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\" @@ -890,9 +887,8 @@ text{*The quotient by the kernel of a homomorphism is isomorphic to the range of that homomorphism.*} -constdefs - kernel :: "('a, 'm) monoid_scheme \ ('b, 'n) monoid_scheme \ - ('a \ 'b) \ 'a set" +definition kernel :: "('a, 'm) monoid_scheme \ ('b, 'n) monoid_scheme \ + ('a \ 'b) \ 'a set" where --{*the kernel of a homomorphism*} "kernel G H h \ {x. x \ carrier G & h x = \\<^bsub>H\<^esub>}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Algebra/Divisibility.thy --- a/src/HOL/Algebra/Divisibility.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Algebra/Divisibility.thy Mon Mar 01 13:42:31 2010 +0100 @@ -3630,8 +3630,7 @@ text {* Number of factors for wellfoundedness *} -constdefs - factorcount :: "_ \ 'a \ nat" +definition factorcount :: "_ \ 'a \ nat" where "factorcount G a == THE c. (ALL as. set as \ carrier G \ wfactors G as a \ c = length as)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Algebra/FiniteProduct.thy --- a/src/HOL/Algebra/FiniteProduct.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Algebra/FiniteProduct.thy Mon Mar 01 13:42:31 2010 +0100 @@ -26,8 +26,7 @@ inductive_cases empty_foldSetDE [elim!]: "({}, x) \ foldSetD D f e" -constdefs - foldD :: "['a set, 'b => 'a => 'a, 'a, 'b set] => 'a" +definition foldD :: "['a set, 'b => 'a => 'a, 'a, 'b set] => 'a" where "foldD D f e A == THE x. (A, x) \ foldSetD D f e" lemma foldSetD_closed: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Algebra/Group.thy --- a/src/HOL/Algebra/Group.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Algebra/Group.thy Mon Mar 01 13:42:31 2010 +0100 @@ -478,8 +478,7 @@ subsection {* Direct Products *} -constdefs - DirProd :: "_ \ _ \ ('a \ 'b) monoid" (infixr "\\" 80) +definition DirProd :: "_ \ _ \ ('a \ 'b) monoid" (infixr "\\" 80) where "G \\ H \ \carrier = carrier G \ carrier H, mult = (\(g, h) (g', h'). (g \\<^bsub>G\<^esub> g', h \\<^bsub>H\<^esub> h')), one = (\\<^bsub>G\<^esub>, \\<^bsub>H\<^esub>)\" @@ -545,8 +544,7 @@ "[|h \ hom G H; i \ hom H I|] ==> compose (carrier G) i h \ hom G I" by (fastsimp simp add: hom_def compose_def) -constdefs - iso :: "_ => _ => ('a => 'b) set" (infixr "\" 60) +definition iso :: "_ => _ => ('a => 'b) set" (infixr "\" 60) where "G \ H == {h. h \ hom G H & bij_betw h (carrier G) (carrier H)}" lemma iso_refl: "(%x. x) \ G \ G" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Algebra/IntRing.thy --- a/src/HOL/Algebra/IntRing.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Algebra/IntRing.thy Mon Mar 01 13:42:31 2010 +0100 @@ -22,8 +22,7 @@ subsection {* @{text "\"}: The Set of Integers as Algebraic Structure *} -constdefs - int_ring :: "int ring" ("\") +definition int_ring :: "int ring" ("\") where "int_ring \ \carrier = UNIV, mult = op *, one = 1, zero = 0, add = op +\" lemma int_Zcarr [intro!, simp]: @@ -324,8 +323,7 @@ subsection {* Ideals and the Modulus *} -constdefs - ZMod :: "int => int => int set" +definition ZMod :: "int => int => int set" where "ZMod k r == (Idl\<^bsub>\\<^esub> {k}) +>\<^bsub>\\<^esub> r" lemmas ZMod_defs = @@ -407,8 +405,7 @@ subsection {* Factorization *} -constdefs - ZFact :: "int \ int set ring" +definition ZFact :: "int \ int set ring" where "ZFact k == \ Quot (Idl\<^bsub>\\<^esub> {k})" lemmas ZFact_defs = ZFact_def FactRing_def diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Algebra/Ring.thy --- a/src/HOL/Algebra/Ring.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Algebra/Ring.thy Mon Mar 01 13:42:31 2010 +0100 @@ -198,8 +198,7 @@ This definition makes it easy to lift lemmas from @{term finprod}. *} -constdefs - finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b" +definition finsum :: "[('b, 'm) ring_scheme, 'a => 'b, 'a set] => 'b" where "finsum G f A == finprod (| carrier = carrier G, mult = add G, one = zero G |) f A" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/CertifiedEmail.thy --- a/src/HOL/Auth/CertifiedEmail.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/CertifiedEmail.thy Mon Mar 01 13:42:31 2010 +0100 @@ -25,8 +25,7 @@ BothAuth :: nat text{*We formalize a fixed way of computing responses. Could be better.*} -constdefs - "response" :: "agent => agent => nat => msg" +definition "response" :: "agent => agent => nat => msg" where "response S R q == Hash {|Agent S, Key (shrK R), Nonce q|}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Guard/Extensions.thy --- a/src/HOL/Auth/Guard/Extensions.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Guard/Extensions.thy Mon Mar 01 13:42:31 2010 +0100 @@ -61,7 +61,7 @@ subsubsection{*messages that are pairs*} -constdefs is_MPair :: "msg => bool" +definition is_MPair :: "msg => bool" where "is_MPair X == EX Y Z. X = {|Y,Z|}" declare is_MPair_def [simp] @@ -96,7 +96,7 @@ declare is_MPair_def [simp del] -constdefs has_no_pair :: "msg set => bool" +definition has_no_pair :: "msg set => bool" where "has_no_pair H == ALL X Y. {|X,Y|} ~:H" declare has_no_pair_def [simp] @@ -117,7 +117,7 @@ subsubsection{*lemmas on keysFor*} -constdefs usekeys :: "msg set => key set" +definition usekeys :: "msg set => key set" where "usekeys G == {K. EX Y. Crypt K Y:G}" lemma finite_keysFor [intro]: "finite G ==> finite (keysFor G)" @@ -237,7 +237,7 @@ subsubsection{*sets of keys*} -constdefs keyset :: "msg set => bool" +definition keyset :: "msg set => bool" where "keyset G == ALL X. X:G --> (EX K. X = Key K)" lemma keyset_in [dest]: "[| keyset G; X:G |] ==> EX K. X = Key K" @@ -257,7 +257,7 @@ subsubsection{*keys a priori necessary for decrypting the messages of G*} -constdefs keysfor :: "msg set => msg set" +definition keysfor :: "msg set => msg set" where "keysfor G == Key ` keysFor (parts G)" lemma keyset_keysfor [iff]: "keyset (keysfor G)" @@ -295,7 +295,7 @@ subsubsection{*general protocol properties*} -constdefs is_Says :: "event => bool" +definition is_Says :: "event => bool" where "is_Says ev == (EX A B X. ev = Says A B X)" lemma is_Says_Says [iff]: "is_Says (Says A B X)" @@ -303,7 +303,7 @@ (* one could also require that Gets occurs after Says but this is sufficient for our purpose *) -constdefs Gets_correct :: "event list set => bool" +definition Gets_correct :: "event list set => bool" where "Gets_correct p == ALL evs B X. evs:p --> Gets B X:set evs --> (EX A. Says A B X:set evs)" @@ -312,7 +312,7 @@ apply (simp add: Gets_correct_def) by (drule_tac x="Gets B X # evs" in spec, auto) -constdefs one_step :: "event list set => bool" +definition one_step :: "event list set => bool" where "one_step p == ALL evs ev. ev#evs:p --> evs:p" lemma one_step_Cons [dest]: "[| one_step p; ev#evs:p |] ==> evs:p" @@ -324,7 +324,7 @@ lemma trunc: "[| evs @ evs':p; one_step p |] ==> evs':p" by (induct evs, auto) -constdefs has_only_Says :: "event list set => bool" +definition has_only_Says :: "event list set => bool" where "has_only_Says p == ALL evs ev. evs:p --> ev:set evs --> (EX A B X. ev = Says A B X)" @@ -450,7 +450,7 @@ if A=A' then insert X (knows_max' A evs) else knows_max' A evs ))" -constdefs knows_max :: "agent => event list => msg set" +definition knows_max :: "agent => event list => msg set" where "knows_max A evs == knows_max' A evs Un initState A" abbreviation @@ -512,7 +512,7 @@ | Notes A X => parts {X} Un used' evs )" -constdefs init :: "msg set" +definition init :: "msg set" where "init == used []" lemma used_decomp: "used evs = init Un used' evs" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Guard/Guard.thy --- a/src/HOL/Auth/Guard/Guard.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Guard/Guard.thy Mon Mar 01 13:42:31 2010 +0100 @@ -76,7 +76,7 @@ subsection{*guarded sets*} -constdefs Guard :: "nat => key set => msg set => bool" +definition Guard :: "nat => key set => msg set => bool" where "Guard n Ks H == ALL X. X:H --> X:guard n Ks" subsection{*basic facts about @{term Guard}*} @@ -241,7 +241,7 @@ subsection{*list corresponding to "decrypt"*} -constdefs decrypt' :: "msg list => key => msg => msg list" +definition decrypt' :: "msg list => key => msg => msg list" where "decrypt' l K Y == Y # remove l (Crypt K Y)" declare decrypt'_def [simp] diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Guard/GuardK.thy --- a/src/HOL/Auth/Guard/GuardK.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Guard/GuardK.thy Mon Mar 01 13:42:31 2010 +0100 @@ -85,7 +85,7 @@ subsection{*guarded sets*} -constdefs GuardK :: "nat => key set => msg set => bool" +definition GuardK :: "nat => key set => msg set => bool" where "GuardK n Ks H == ALL X. X:H --> X:guardK n Ks" subsection{*basic facts about @{term GuardK}*} @@ -239,7 +239,7 @@ subsection{*list corresponding to "decrypt"*} -constdefs decrypt' :: "msg list => key => msg => msg list" +definition decrypt' :: "msg list => key => msg => msg list" where "decrypt' l K Y == Y # remove l (Crypt K Y)" declare decrypt'_def [simp] diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Guard/Guard_Public.thy --- a/src/HOL/Auth/Guard/Guard_Public.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Guard/Guard_Public.thy Mon Mar 01 13:42:31 2010 +0100 @@ -19,7 +19,7 @@ subsubsection{*signature*} -constdefs sign :: "agent => msg => msg" +definition sign :: "agent => msg => msg" where "sign A X == {|Agent A, X, Crypt (priK A) (Hash X)|}" lemma sign_inj [iff]: "(sign A X = sign A' X') = (A=A' & X=X')" @@ -27,7 +27,7 @@ subsubsection{*agent associated to a key*} -constdefs agt :: "key => agent" +definition agt :: "key => agent" where "agt K == @A. K = priK A | K = pubK A" lemma agt_priK [simp]: "agt (priK A) = A" @@ -57,7 +57,7 @@ subsubsection{*sets of private keys*} -constdefs priK_set :: "key set => bool" +definition priK_set :: "key set => bool" where "priK_set Ks == ALL K. K:Ks --> (EX A. K = priK A)" lemma in_priK_set: "[| priK_set Ks; K:Ks |] ==> EX A. K = priK A" @@ -71,7 +71,7 @@ subsubsection{*sets of good keys*} -constdefs good :: "key set => bool" +definition good :: "key set => bool" where "good Ks == ALL K. K:Ks --> agt K ~:bad" lemma in_good: "[| good Ks; K:Ks |] ==> agt K ~:bad" @@ -99,7 +99,7 @@ subsubsection{*function giving a new nonce*} -constdefs new :: "event list => nat" +definition new :: "event list => nat" where "new evs == Suc (greatest evs)" lemma new_isnt_used [iff]: "Nonce (new evs) ~:used evs" @@ -151,7 +151,7 @@ subsubsection{*regular protocols*} -constdefs regular :: "event list set => bool" +definition regular :: "event list set => bool" where "regular p == ALL evs A. evs:p --> (Key (priK A):parts (spies evs)) = (A:bad)" lemma priK_parts_iff_bad [simp]: "[| evs:p; regular p |] ==> diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Guard/Guard_Shared.thy --- a/src/HOL/Auth/Guard/Guard_Shared.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Guard/Guard_Shared.thy Mon Mar 01 13:42:31 2010 +0100 @@ -25,7 +25,7 @@ subsubsection{*agent associated to a key*} -constdefs agt :: "key => agent" +definition agt :: "key => agent" where "agt K == @A. K = shrK A" lemma agt_shrK [simp]: "agt (shrK A) = A" @@ -52,7 +52,7 @@ subsubsection{*sets of symmetric keys*} -constdefs shrK_set :: "key set => bool" +definition shrK_set :: "key set => bool" where "shrK_set Ks == ALL K. K:Ks --> (EX A. K = shrK A)" lemma in_shrK_set: "[| shrK_set Ks; K:Ks |] ==> EX A. K = shrK A" @@ -66,7 +66,7 @@ subsubsection{*sets of good keys*} -constdefs good :: "key set => bool" +definition good :: "key set => bool" where "good Ks == ALL K. K:Ks --> agt K ~:bad" lemma in_good: "[| good Ks; K:Ks |] ==> agt K ~:bad" @@ -154,7 +154,7 @@ subsubsection{*regular protocols*} -constdefs regular :: "event list set => bool" +definition regular :: "event list set => bool" where "regular p == ALL evs A. evs:p --> (Key (shrK A):parts (spies evs)) = (A:bad)" lemma shrK_parts_iff_bad [simp]: "[| evs:p; regular p |] ==> diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Guard/Guard_Yahalom.thy --- a/src/HOL/Auth/Guard/Guard_Yahalom.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy Mon Mar 01 13:42:31 2010 +0100 @@ -198,7 +198,7 @@ subsection{*guardedness of NB*} -constdefs ya_keys :: "agent => agent => nat => nat => event list => key set" +definition ya_keys :: "agent => agent => nat => nat => event list => key set" where "ya_keys A B NA NB evs == {shrK A,shrK B} Un {K. ya3 A B NA NB K:set evs}" lemma Guard_NB [rule_format]: "[| evs:ya; A ~:bad; B ~:bad |] ==> diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Guard/P1.thy --- a/src/HOL/Auth/Guard/P1.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Guard/P1.thy Mon Mar 01 13:42:31 2010 +0100 @@ -39,7 +39,7 @@ subsubsection{*offer chaining: B chains his offer for A with the head offer of L for sending it to C*} -constdefs chain :: "agent => nat => agent => msg => agent => msg" +definition chain :: "agent => nat => agent => msg => agent => msg" where "chain B ofr A L C == let m1= Crypt (pubK A) (Nonce ofr) in let m2= Hash {|head L, Agent C|} in @@ -86,7 +86,7 @@ subsubsection{*anchor of the offer list*} -constdefs anchor :: "agent => nat => agent => msg" +definition anchor :: "agent => nat => agent => msg" where "anchor A n B == chain A n A (cons nil nil) B" lemma anchor_inj [iff]: "(anchor A n B = anchor A' n' B') @@ -107,7 +107,7 @@ subsubsection{*request event*} -constdefs reqm :: "agent => nat => nat => msg => agent => msg" +definition reqm :: "agent => nat => nat => msg => agent => msg" where "reqm A r n I B == {|Agent A, Number r, cons (Agent A) (cons (Agent B) I), cons (anchor A n B) nil|}" @@ -118,7 +118,7 @@ lemma Nonce_in_reqm [iff]: "Nonce n:parts {reqm A r n I B}" by (auto simp: reqm_def) -constdefs req :: "agent => nat => nat => msg => agent => event" +definition req :: "agent => nat => nat => msg => agent => event" where "req A r n I B == Says A B (reqm A r n I B)" lemma req_inj [iff]: "(req A r n I B = req A' r' n' I' B') @@ -127,8 +127,8 @@ subsubsection{*propose event*} -constdefs prom :: "agent => nat => agent => nat => msg => msg => -msg => agent => msg" +definition prom :: "agent => nat => agent => nat => msg => msg => +msg => agent => msg" where "prom B ofr A r I L J C == {|Agent A, Number r, app (J, del (Agent B, I)), cons (chain B ofr A L C) L|}" @@ -140,8 +140,8 @@ lemma Nonce_in_prom [iff]: "Nonce ofr:parts {prom B ofr A r I L J C}" by (auto simp: prom_def) -constdefs pro :: "agent => nat => agent => nat => msg => msg => -msg => agent => event" +definition pro :: "agent => nat => agent => nat => msg => msg => +msg => agent => event" where "pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)" lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C' @@ -198,7 +198,7 @@ subsubsection{*offers of an offer list*} -constdefs offer_nonces :: "msg => msg set" +definition offer_nonces :: "msg => msg set" where "offer_nonces L == {X. X:parts {L} & (EX n. X = Nonce n)}" subsubsection{*the originator can get the offers*} @@ -252,7 +252,7 @@ pro C (Suc ofr) A r I' L nil D # trace (B,Suc ofr,A,r,I'',tail L,K))" -constdefs trace' :: "agent => nat => nat => msg => agent => nat => event list" +definition trace' :: "agent => nat => nat => msg => agent => nat => event list" where "trace' A r n I B ofr == ( let AI = cons (Agent A) I in let L = offer_list (A,n,B,AI,ofr) in diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Guard/P2.thy --- a/src/HOL/Auth/Guard/P2.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Guard/P2.thy Mon Mar 01 13:42:31 2010 +0100 @@ -26,7 +26,7 @@ subsubsection{*offer chaining: B chains his offer for A with the head offer of L for sending it to C*} -constdefs chain :: "agent => nat => agent => msg => agent => msg" +definition chain :: "agent => nat => agent => msg => agent => msg" where "chain B ofr A L C == let m1= sign B (Nonce ofr) in let m2= Hash {|head L, Agent C|} in @@ -73,7 +73,7 @@ subsubsection{*anchor of the offer list*} -constdefs anchor :: "agent => nat => agent => msg" +definition anchor :: "agent => nat => agent => msg" where "anchor A n B == chain A n A (cons nil nil) B" lemma anchor_inj [iff]: @@ -88,7 +88,7 @@ subsubsection{*request event*} -constdefs reqm :: "agent => nat => nat => msg => agent => msg" +definition reqm :: "agent => nat => nat => msg => agent => msg" where "reqm A r n I B == {|Agent A, Number r, cons (Agent A) (cons (Agent B) I), cons (anchor A n B) nil|}" @@ -99,7 +99,7 @@ lemma Nonce_in_reqm [iff]: "Nonce n:parts {reqm A r n I B}" by (auto simp: reqm_def) -constdefs req :: "agent => nat => nat => msg => agent => event" +definition req :: "agent => nat => nat => msg => agent => event" where "req A r n I B == Says A B (reqm A r n I B)" lemma req_inj [iff]: "(req A r n I B = req A' r' n' I' B') @@ -108,8 +108,8 @@ subsubsection{*propose event*} -constdefs prom :: "agent => nat => agent => nat => msg => msg => -msg => agent => msg" +definition prom :: "agent => nat => agent => nat => msg => msg => +msg => agent => msg" where "prom B ofr A r I L J C == {|Agent A, Number r, app (J, del (Agent B, I)), cons (chain B ofr A L C) L|}" @@ -120,8 +120,8 @@ lemma Nonce_in_prom [iff]: "Nonce ofr:parts {prom B ofr A r I L J C}" by (auto simp: prom_def) -constdefs pro :: "agent => nat => agent => nat => msg => msg => - msg => agent => event" +definition pro :: "agent => nat => agent => nat => msg => msg => + msg => agent => event" where "pro B ofr A r I L J C == Says B C (prom B ofr A r I L J C)" lemma pro_inj [dest]: "pro B ofr A r I L J C = pro B' ofr' A' r' I' L' J' C' diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Guard/Proto.thy --- a/src/HOL/Auth/Guard/Proto.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Guard/Proto.thy Mon Mar 01 13:42:31 2010 +0100 @@ -23,7 +23,7 @@ types proto = "rule set" -constdefs wdef :: "proto => bool" +definition wdef :: "proto => bool" where "wdef p == ALL R k. R:p --> Number k:parts {msg' R} --> Number k:parts (msg`(fst R))" @@ -35,19 +35,17 @@ nb :: "nat => msg" key :: "key => key" -consts apm :: "subs => msg => msg" - -primrec -"apm s (Agent A) = Agent (agent s A)" -"apm s (Nonce n) = Nonce (nonce s n)" -"apm s (Number n) = nb s n" -"apm s (Key K) = Key (key s K)" -"apm s (Hash X) = Hash (apm s X)" -"apm s (Crypt K X) = ( +primrec apm :: "subs => msg => msg" where + "apm s (Agent A) = Agent (agent s A)" +| "apm s (Nonce n) = Nonce (nonce s n)" +| "apm s (Number n) = nb s n" +| "apm s (Key K) = Key (key s K)" +| "apm s (Hash X) = Hash (apm s X)" +| "apm s (Crypt K X) = ( if (EX A. K = pubK A) then Crypt (pubK (agent s (agt K))) (apm s X) else if (EX A. K = priK A) then Crypt (priK (agent s (agt K))) (apm s X) else Crypt (key s K) (apm s X))" -"apm s {|X,Y|} = {|apm s X, apm s Y|}" +| "apm s {|X,Y|} = {|apm s X, apm s Y|}" lemma apm_parts: "X:parts {Y} ==> apm s X:parts {apm s Y}" apply (erule parts.induct, simp_all, blast) @@ -69,12 +67,10 @@ apply (drule_tac Y="msg x" and s=s in apm_parts, simp) by (blast dest: parts_parts) -consts ap :: "subs => event => event" - -primrec -"ap s (Says A B X) = Says (agent s A) (agent s B) (apm s X)" -"ap s (Gets A X) = Gets (agent s A) (apm s X)" -"ap s (Notes A X) = Notes (agent s A) (apm s X)" +primrec ap :: "subs => event => event" where + "ap s (Says A B X) = Says (agent s A) (agent s B) (apm s X)" +| "ap s (Gets A X) = Gets (agent s A) (apm s X)" +| "ap s (Notes A X) = Notes (agent s A) (apm s X)" abbreviation ap' :: "subs => rule => event" where @@ -94,7 +90,7 @@ subsection{*nonces generated by a rule*} -constdefs newn :: "rule => nat set" +definition newn :: "rule => nat set" where "newn R == {n. Nonce n:parts {msg (snd R)} & Nonce n ~:parts (msg`(fst R))}" lemma newn_parts: "n:newn R ==> Nonce (nonce s n):parts {apm' s R}" @@ -102,7 +98,7 @@ subsection{*traces generated by a protocol*} -constdefs ok :: "event list => rule => subs => bool" +definition ok :: "event list => rule => subs => bool" where "ok evs R s == ((ALL x. x:fst R --> ap s x:set evs) & (ALL n. n:newn R --> Nonce (nonce s n) ~:used evs))" @@ -124,7 +120,7 @@ apply (unfold one_step_def, clarify) by (ind_cases "ev # evs:tr p" for ev evs, auto) -constdefs has_only_Says' :: "proto => bool" +definition has_only_Says' :: "proto => bool" where "has_only_Says' p == ALL R. R:p --> is_Says (snd R)" lemma has_only_Says'D: "[| R:p; has_only_Says' p |] @@ -165,8 +161,8 @@ subsection{*introduction of a fresh guarded nonce*} -constdefs fresh :: "proto => rule => subs => nat => key set => event list -=> bool" +definition fresh :: "proto => rule => subs => nat => key set => event list +=> bool" where "fresh p R s n Ks evs == (EX evs1 evs2. evs = evs2 @ ap' s R # evs1 & Nonce n ~:used evs1 & R:p & ok evs1 R s & Nonce n:parts {apm' s R} & apm' s R:guard n Ks)" @@ -226,7 +222,7 @@ subsection{*safe keys*} -constdefs safe :: "key set => msg set => bool" +definition safe :: "key set => msg set => bool" where "safe Ks G == ALL K. K:Ks --> Key K ~:analz G" lemma safeD [dest]: "[| safe Ks G; K:Ks |] ==> Key K ~:analz G" @@ -240,7 +236,7 @@ subsection{*guardedness preservation*} -constdefs preserv :: "proto => keyfun => nat => key set => bool" +definition preserv :: "proto => keyfun => nat => key set => bool" where "preserv p keys n Ks == (ALL evs R' s' R s. evs:tr p --> Guard n Ks (spies evs) --> safe Ks (spies evs) --> fresh p R' s' n Ks evs --> keys R' s' n evs <= Ks --> R:p --> ok evs R s --> apm' s R:guard n Ks)" @@ -257,7 +253,7 @@ subsection{*monotonic keyfun*} -constdefs monoton :: "proto => keyfun => bool" +definition monoton :: "proto => keyfun => bool" where "monoton p keys == ALL R' s' n ev evs. ev # evs:tr p --> keys R' s' n evs <= keys R' s' n (ev # evs)" @@ -323,7 +319,7 @@ subsection{*unicity*} -constdefs uniq :: "proto => secfun => bool" +definition uniq :: "proto => secfun => bool" where "uniq p secret == ALL evs R R' n n' Ks s s'. R:p --> R':p --> n:newn R --> n':newn R' --> nonce s n = nonce s' n' --> Nonce (nonce s n):parts {apm' s R} --> Nonce (nonce s n):parts {apm' s' R'} --> @@ -340,13 +336,13 @@ secret R n s Ks = secret R' n' s' Ks" by (unfold uniq_def, blast) -constdefs ord :: "proto => (rule => rule => bool) => bool" +definition ord :: "proto => (rule => rule => bool) => bool" where "ord p inff == ALL R R'. R:p --> R':p --> ~ inff R R' --> inff R' R" lemma ordD: "[| ord p inff; ~ inff R R'; R:p; R':p |] ==> inff R' R" by (unfold ord_def, blast) -constdefs uniq' :: "proto => (rule => rule => bool) => secfun => bool" +definition uniq' :: "proto => (rule => rule => bool) => secfun => bool" where "uniq' p inff secret == ALL evs R R' n n' Ks s s'. R:p --> R':p --> inff R R' --> n:newn R --> n':newn R' --> nonce s n = nonce s' n' --> Nonce (nonce s n):parts {apm' s R} --> Nonce (nonce s n):parts {apm' s' R'} --> @@ -372,13 +368,12 @@ subsection{*Needham-Schroeder-Lowe*} -constdefs -a :: agent "a == Friend 0" -b :: agent "b == Friend 1" -a' :: agent "a' == Friend 2" -b' :: agent "b' == Friend 3" -Na :: nat "Na == 0" -Nb :: nat "Nb == 1" +definition a :: agent where "a == Friend 0" +definition b :: agent where "b == Friend 1" +definition a' :: agent where "a' == Friend 2" +definition b' :: agent where "b' == Friend 3" +definition Na :: nat where "Na == 0" +definition Nb :: nat where "Nb == 1" abbreviation ns1 :: rule where @@ -408,19 +403,19 @@ ns3b :: event where "ns3b == Says b' a (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|})" -constdefs keys :: "keyfun" +definition keys :: "keyfun" where "keys R' s' n evs == {priK' s' a, priK' s' b}" lemma "monoton ns keys" by (simp add: keys_def monoton_def) -constdefs secret :: "secfun" +definition secret :: "secfun" where "secret R n s Ks == (if R=ns1 then apm s (Crypt (pubK b) {|Nonce Na, Agent a|}) else if R=ns2 then apm s (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|}) else Number 0)" -constdefs inf :: "rule => rule => bool" +definition inf :: "rule => rule => bool" where "inf R R' == (R=ns1 | (R=ns2 & R'~=ns1) | (R=ns3 & R'=ns3))" lemma inf_is_ord [iff]: "ord ns inf" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/KerberosIV.thy --- a/src/HOL/Auth/KerberosIV.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/KerberosIV.thy Mon Mar 01 13:42:31 2010 +0100 @@ -101,8 +101,7 @@ (* Predicate formalising the association between authKeys and servKeys *) -constdefs - AKcryptSK :: "[key, key, event list] => bool" +definition AKcryptSK :: "[key, key, event list] => bool" where "AKcryptSK authK servK evs == \A B Ts. Says Tgs A (Crypt authK diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/KerberosIV_Gets.thy --- a/src/HOL/Auth/KerberosIV_Gets.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/KerberosIV_Gets.thy Mon Mar 01 13:42:31 2010 +0100 @@ -89,8 +89,7 @@ (* Predicate formalising the association between authKeys and servKeys *) -constdefs - AKcryptSK :: "[key, key, event list] => bool" +definition AKcryptSK :: "[key, key, event list] => bool" where "AKcryptSK authK servK evs == \A B Ts. Says Tgs A (Crypt authK diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/KerberosV.thy --- a/src/HOL/Auth/KerberosV.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/KerberosV.thy Mon Mar 01 13:42:31 2010 +0100 @@ -92,8 +92,7 @@ (* Predicate formalising the association between authKeys and servKeys *) -constdefs - AKcryptSK :: "[key, key, event list] => bool" +definition AKcryptSK :: "[key, key, event list] => bool" where "AKcryptSK authK servK evs == \A B tt. Says Tgs A \Crypt authK \Key servK, Agent B, tt\, diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Message.thy --- a/src/HOL/Auth/Message.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Message.thy Mon Mar 01 13:42:31 2010 +0100 @@ -32,8 +32,7 @@ text{*The inverse of a symmetric key is itself; that of a public key is the private key and vice versa*} -constdefs - symKeys :: "key set" +definition symKeys :: "key set" where "symKeys == {K. invKey K = K}" datatype --{*We allow any number of friendly agents*} @@ -61,12 +60,11 @@ "{|x, y|}" == "CONST MPair x y" -constdefs - HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) +definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where --{*Message Y paired with a MAC computed with the help of X*} "Hash[X] Y == {| Hash{|X,Y|}, Y|}" - keysFor :: "msg set => key set" +definition keysFor :: "msg set => key set" where --{*Keys useful to decrypt elements of a message set*} "keysFor H == invKey ` {K. \X. Crypt K X \ H}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Smartcard/ShoupRubin.thy --- a/src/HOL/Auth/Smartcard/ShoupRubin.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ -(* ID: $Id$ - Author: Giampaolo Bella, Catania University +(* Author: Giampaolo Bella, Catania University *) header{*Original Shoup-Rubin protocol*} @@ -29,9 +28,7 @@ between each agent and his smartcard*) shouprubin_assumes_securemeans [iff]: "evs \ sr \ secureM" -constdefs - - Unique :: "[event, event list] => bool" ("Unique _ on _") +definition Unique :: "[event, event list] => bool" ("Unique _ on _") where "Unique ev on evs == ev \ set (tl (dropWhile (% z. z \ ev) evs))" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Smartcard/ShoupRubinBella.thy --- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ -(* ID: $Id$ - Author: Giampaolo Bella, Catania University +(* Author: Giampaolo Bella, Catania University *) header{*Bella's modification of the Shoup-Rubin protocol*} @@ -35,9 +34,7 @@ between each agent and his smartcard*) shouprubin_assumes_securemeans [iff]: "evs \ srb \ secureM" -constdefs - - Unique :: "[event, event list] => bool" ("Unique _ on _") +definition Unique :: "[event, event list] => bool" ("Unique _ on _") where "Unique ev on evs == ev \ set (tl (dropWhile (% z. z \ ev) evs))" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Smartcard/Smartcard.thy --- a/src/HOL/Auth/Smartcard/Smartcard.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Mon Mar 01 13:42:31 2010 +0100 @@ -43,15 +43,11 @@ shrK_disj_pin [iff]: "shrK P \ pin Q" crdK_disj_pin [iff]: "crdK C \ pin P" -constdefs - legalUse :: "card => bool" ("legalUse (_)") +definition legalUse :: "card => bool" ("legalUse (_)") where "legalUse C == C \ stolen" -consts - illegalUse :: "card => bool" -primrec - illegalUse_def: - "illegalUse (Card A) = ( (Card A \ stolen \ A \ bad) \ Card A \ cloned )" +primrec illegalUse :: "card => bool" where + illegalUse_def: "illegalUse (Card A) = ( (Card A \ stolen \ A \ bad) \ Card A \ cloned )" text{*initState must be defined with care*} diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/TLS.thy --- a/src/HOL/Auth/TLS.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/TLS.thy Mon Mar 01 13:42:31 2010 +0100 @@ -43,8 +43,7 @@ theory TLS imports Public Nat_Int_Bij begin -constdefs - certificate :: "[agent,key] => msg" +definition certificate :: "[agent,key] => msg" where "certificate A KA == Crypt (priSK Server) {|Agent A, Key KA|}" text{*TLS apparently does not require separate keypairs for encryption and diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/Yahalom.thy --- a/src/HOL/Auth/Yahalom.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/Yahalom.thy Mon Mar 01 13:42:31 2010 +0100 @@ -74,8 +74,7 @@ ==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \ yahalom" -constdefs - KeyWithNonce :: "[key, nat, event list] => bool" +definition KeyWithNonce :: "[key, nat, event list] => bool" where "KeyWithNonce K NB evs == \A B na X. Says Server A {|Crypt (shrK A) {|Agent B, Key K, na, Nonce NB|}, X|} diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Auth/ZhouGollmann.thy --- a/src/HOL/Auth/ZhouGollmann.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Auth/ZhouGollmann.thy Mon Mar 01 13:42:31 2010 +0100 @@ -21,8 +21,7 @@ abbreviation f_con :: nat where "f_con == 4" -constdefs - broken :: "agent set" +definition broken :: "agent set" where --{*the compromised honest agents; TTP is included as it's not allowed to use the protocol*} "broken == bad - {Spy}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/AxCompl.thy --- a/src/HOL/Bali/AxCompl.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/AxCompl.thy Mon Mar 01 13:42:31 2010 +0100 @@ -20,9 +20,7 @@ section "set of not yet initialzed classes" -constdefs - - nyinitcls :: "prog \ state \ qtname set" +definition nyinitcls :: "prog \ state \ qtname set" where "nyinitcls G s \ {C. is_class G C \ \ initd C s}" lemma nyinitcls_subset_class: "nyinitcls G s \ {C. is_class G C}" @@ -115,8 +113,7 @@ section "init-le" -constdefs - init_le :: "prog \ nat \ state \ bool" ("_\init\_" [51,51] 50) +definition init_le :: "prog \ nat \ state \ bool" ("_\init\_" [51,51] 50) where "G\init\n \ \s. card (nyinitcls G s) \ n" lemma init_le_def2 [simp]: "(G\init\n) s = (card (nyinitcls G s)\n)" @@ -135,9 +132,7 @@ section "Most General Triples and Formulas" -constdefs - - remember_init_state :: "state assn" ("\") +definition remember_init_state :: "state assn" ("\") where "\ \ \Y s Z. s = Z" lemma remember_init_state_def2 [simp]: "\ Y = op =" @@ -579,8 +574,7 @@ unroll the loop in iterated evaluations of the expression and evaluations of the loop body. *} -constdefs - unroll:: "prog \ label \ expr \ stmt \ (state \ state) set" +definition unroll :: "prog \ label \ expr \ stmt \ (state \ state) set" where "unroll G l e c \ {(s,t). \ v s1 s2. G\s \e-\v\ s1 \ the_Bool v \ normal s1 \ diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/AxExample.thy --- a/src/HOL/Bali/AxExample.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/AxExample.thy Mon Mar 01 13:42:31 2010 +0100 @@ -8,8 +8,7 @@ imports AxSem Example begin -constdefs - arr_inv :: "st \ bool" +definition arr_inv :: "st \ bool" where "arr_inv \ \s. \obj a T el. globs s (Stat Base) = Some obj \ values obj (Inl (arr, Base)) = Some (Addr a) \ heap s a = Some \tag=Arr T 2,values=el\" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/AxSem.thy --- a/src/HOL/Bali/AxSem.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/AxSem.thy Mon Mar 01 13:42:31 2010 +0100 @@ -63,8 +63,7 @@ "res" <= (type) "AxSem.res" "a assn" <= (type) "vals \ state \ a \ bool" -constdefs - assn_imp :: "'a assn \ 'a assn \ bool" (infixr "\" 25) +definition assn_imp :: "'a assn \ 'a assn \ bool" (infixr "\" 25) where "P \ Q \ \Y s Z. P Y s Z \ Q Y s Z" lemma assn_imp_def2 [iff]: "(P \ Q) = (\Y s Z. P Y s Z \ Q Y s Z)" @@ -77,8 +76,7 @@ subsection "peek-and" -constdefs - peek_and :: "'a assn \ (state \ bool) \ 'a assn" (infixl "\." 13) +definition peek_and :: "'a assn \ (state \ bool) \ 'a assn" (infixl "\." 13) where "P \. p \ \Y s Z. P Y s Z \ p s" lemma peek_and_def2 [simp]: "peek_and P p Y s = (\Z. (P Y s Z \ p s))" @@ -117,8 +115,7 @@ subsection "assn-supd" -constdefs - assn_supd :: "'a assn \ (state \ state) \ 'a assn" (infixl ";." 13) +definition assn_supd :: "'a assn \ (state \ state) \ 'a assn" (infixl ";." 13) where "P ;. f \ \Y s' Z. \s. P Y s Z \ s' = f s" lemma assn_supd_def2 [simp]: "assn_supd P f Y s' Z = (\s. P Y s Z \ s' = f s)" @@ -128,8 +125,7 @@ subsection "supd-assn" -constdefs - supd_assn :: "(state \ state) \ 'a assn \ 'a assn" (infixr ".;" 13) +definition supd_assn :: "(state \ state) \ 'a assn \ 'a assn" (infixr ".;" 13) where "f .; P \ \Y s. P Y (f s)" @@ -148,8 +144,7 @@ subsection "subst-res" -constdefs - subst_res :: "'a assn \ res \ 'a assn" ("_\_" [60,61] 60) +definition subst_res :: "'a assn \ res \ 'a assn" ("_\_" [60,61] 60) where "P\w \ \Y. P w" lemma subst_res_def2 [simp]: "(P\w) Y = P w" @@ -184,8 +179,7 @@ subsection "subst-Bool" -constdefs - subst_Bool :: "'a assn \ bool \ 'a assn" ("_\=_" [60,61] 60) +definition subst_Bool :: "'a assn \ bool \ 'a assn" ("_\=_" [60,61] 60) where "P\=b \ \Y s Z. \v. P (Val v) s Z \ (normal s \ the_Bool v=b)" lemma subst_Bool_def2 [simp]: @@ -200,8 +194,7 @@ subsection "peek-res" -constdefs - peek_res :: "(res \ 'a assn) \ 'a assn" +definition peek_res :: "(res \ 'a assn) \ 'a assn" where "peek_res Pf \ \Y. Pf Y Y" syntax @@ -229,8 +222,7 @@ subsection "ign-res" -constdefs - ign_res :: " 'a assn \ 'a assn" ("_\" [1000] 1000) +definition ign_res :: " 'a assn \ 'a assn" ("_\" [1000] 1000) where "P\ \ \Y s Z. \Y. P Y s Z" lemma ign_res_def2 [simp]: "P\ Y s Z = (\Y. P Y s Z)" @@ -261,8 +253,7 @@ subsection "peek-st" -constdefs - peek_st :: "(st \ 'a assn) \ 'a assn" +definition peek_st :: "(st \ 'a assn) \ 'a assn" where "peek_st P \ \Y s. P (store s) Y s" syntax @@ -306,8 +297,7 @@ subsection "ign-res-eq" -constdefs - ign_res_eq :: "'a assn \ res \ 'a assn" ("_\=_" [60,61] 60) +definition ign_res_eq :: "'a assn \ res \ 'a assn" ("_\=_" [60,61] 60) where "P\=w \ \Y:. P\ \. (\s. Y=w)" lemma ign_res_eq_def2 [simp]: "(P\=w) Y s Z = ((\Y. P Y s Z) \ Y=w)" @@ -337,8 +327,7 @@ subsection "RefVar" -constdefs - RefVar :: "(state \ vvar \ state) \ 'a assn \ 'a assn"(infixr "..;" 13) +definition RefVar :: "(state \ vvar \ state) \ 'a assn \ 'a assn" (infixr "..;" 13) where "vf ..; P \ \Y s. let (v,s') = vf s in P (Var v) s'" lemma RefVar_def2 [simp]: "(vf ..; P) Y s = @@ -349,12 +338,11 @@ subsection "allocation" -constdefs - Alloc :: "prog \ obj_tag \ 'a assn \ 'a assn" +definition Alloc :: "prog \ obj_tag \ 'a assn \ 'a assn" where "Alloc G otag P \ \Y s Z. \s' a. G\s \halloc otag\a\ s'\ P (Val (Addr a)) s' Z" - SXAlloc :: "prog \ 'a assn \ 'a assn" +definition SXAlloc :: "prog \ 'a assn \ 'a assn" where "SXAlloc G P \ \Y s Z. \s'. G\s \sxalloc\ s' \ P Y s' Z" @@ -372,8 +360,7 @@ section "validity" -constdefs - type_ok :: "prog \ term \ state \ bool" +definition type_ok :: "prog \ term \ state \ bool" where "type_ok G t s \ \L T C A. (normal s \ \prg=G,cls=C,lcl=L\\t\T \ \prg=G,cls=C,lcl=L\\dom (locals (store s))\t\A ) @@ -419,10 +406,8 @@ apply auto done -constdefs - mtriples :: "('c \ 'sig \ 'a assn) \ ('c \ 'sig \ expr) \ - ('c \ 'sig \ 'a assn) \ ('c \ 'sig) set \ 'a triples" - ("{{(1_)}/ _-\/ {(1_)} | _}"[3,65,3,65]75) +definition mtriples :: "('c \ 'sig \ 'a assn) \ ('c \ 'sig \ expr) \ + ('c \ 'sig \ 'a assn) \ ('c \ 'sig) set \ 'a triples" ("{{(1_)}/ _-\/ {(1_)} | _}"[3,65,3,65]75) where "{{P} tf-\ {Q} | ms} \ (\(C,sig). {Normal(P C sig)} tf C sig-\ {Q C sig})`ms" consts @@ -641,8 +626,7 @@ axioms *) -constdefs - adapt_pre :: "'a assn \ 'a assn \ 'a assn \ 'a assn" +definition adapt_pre :: "'a assn \ 'a assn \ 'a assn \ 'a assn" where "adapt_pre P Q Q'\\Y s Z. \Y' s'. \Z'. P Y s Z' \ (Q Y' s' Z' \ Q' Y' s' Z)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/Basis.thy --- a/src/HOL/Bali/Basis.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/Basis.thy Mon Mar 01 13:42:31 2010 +0100 @@ -237,8 +237,7 @@ text{* Deemed too special for theory Map. *} -constdefs - chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)" +definition chg_map :: "('b => 'b) => 'a => ('a ~=> 'b) => ('a ~=> 'b)" where "chg_map f a m == case m a of None => m | Some b => m(a|->f b)" lemma chg_map_new[simp]: "m a = None ==> chg_map f a m = m" @@ -253,8 +252,7 @@ section "unique association lists" -constdefs - unique :: "('a \ 'b) list \ bool" +definition unique :: "('a \ 'b) list \ bool" where "unique \ distinct \ map fst" lemma uniqueD [rule_format (no_asm)]: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/Conform.thy --- a/src/HOL/Bali/Conform.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/Conform.thy Mon Mar 01 13:42:31 2010 +0100 @@ -22,9 +22,7 @@ section "extension of global store" -constdefs - - gext :: "st \ st \ bool" ("_\|_" [71,71] 70) +definition gext :: "st \ st \ bool" ("_\|_" [71,71] 70) where "s\|s' \ \r. \obj\globs s r: \obj'\globs s' r: tag obj'= tag obj" text {* For the the proof of type soundness we will need the @@ -98,9 +96,7 @@ section "value conformance" -constdefs - - conf :: "prog \ st \ val \ ty \ bool" ("_,_\_\\_" [71,71,71,71] 70) +definition conf :: "prog \ st \ val \ ty \ bool" ("_,_\_\\_" [71,71,71,71] 70) where "G,s\v\\T \ \T'\typeof (\a. Option.map obj_ty (heap s a)) v:G\T'\T" lemma conf_cong [simp]: "G,set_locals l s\v\\T = G,s\v\\T" @@ -181,10 +177,7 @@ section "value list conformance" -constdefs - - lconf :: "prog \ st \ ('a, val) table \ ('a, ty) table \ bool" - ("_,_\_[\\]_" [71,71,71,71] 70) +definition lconf :: "prog \ st \ ('a, val) table \ ('a, ty) table \ bool" ("_,_\_[\\]_" [71,71,71,71] 70) where "G,s\vs[\\]Ts \ \n. \T\Ts n: \v\vs n: G,s\v\\T" lemma lconfD: "\G,s\vs[\\]Ts; Ts n = Some T\ \ G,s\(the (vs n))\\T" @@ -267,10 +260,7 @@ *} -constdefs - - wlconf :: "prog \ st \ ('a, val) table \ ('a, ty) table \ bool" - ("_,_\_[\\\]_" [71,71,71,71] 70) +definition wlconf :: "prog \ st \ ('a, val) table \ ('a, ty) table \ bool" ("_,_\_[\\\]_" [71,71,71,71] 70) where "G,s\vs[\\\]Ts \ \n. \T\Ts n: \ v\vs n: G,s\v\\T" lemma wlconfD: "\G,s\vs[\\\]Ts; Ts n = Some T; vs n = Some v\ \ G,s\v\\T" @@ -348,9 +338,7 @@ section "object conformance" -constdefs - - oconf :: "prog \ st \ obj \ oref \ bool" ("_,_\_\\\_" [71,71,71,71] 70) +definition oconf :: "prog \ st \ obj \ oref \ bool" ("_,_\_\\\_" [71,71,71,71] 70) where "G,s\obj\\\r \ G,s\values obj[\\]var_tys G (tag obj) r \ (case r of Heap a \ is_type G (obj_ty obj) @@ -385,9 +373,7 @@ section "state conformance" -constdefs - - conforms :: "state \ env' \ bool" ( "_\\_" [71,71] 70) +definition conforms :: "state \ env' \ bool" ("_\\_" [71,71] 70) where "xs\\E \ let (G, L) = E; s = snd xs; l = locals s in (\r. \obj\globs s r: G,s\obj \\\r) \ \ G,s\l [\\\]L\ \ diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/Decl.thy --- a/src/HOL/Bali/Decl.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/Decl.thy Mon Mar 01 13:42:31 2010 +0100 @@ -206,8 +206,7 @@ "mdecl" <= (type) "sig \ methd" -constdefs - mhead::"methd \ mhead" +definition mhead :: "methd \ mhead" where "mhead m \ \access=access m, static=static m, pars=pars m, resT=resT m\" lemma access_mhead [simp]:"access (mhead m) = access m" @@ -275,7 +274,7 @@ lemma memberid_pair_simp1: "memberid p = memberid (snd p)" by (simp add: pair_memberid_def) -constdefs is_field :: "qtname \ memberdecl \ bool" +definition is_field :: "qtname \ memberdecl \ bool" where "is_field m \ \ declC f. m=(declC,fdecl f)" lemma is_fieldD: "is_field m \ \ declC f. m=(declC,fdecl f)" @@ -284,7 +283,7 @@ lemma is_fieldI: "is_field (C,fdecl f)" by (simp add: is_field_def) -constdefs is_method :: "qtname \ memberdecl \ bool" +definition is_method :: "qtname \ memberdecl \ bool" where "is_method membr \ \ declC m. membr=(declC,mdecl m)" lemma is_methodD: "is_method membr \ \ declC m. membr=(declC,mdecl m)" @@ -315,8 +314,7 @@ isuperIfs::qtname list,\::'a\" "idecl" <= (type) "qtname \ iface" -constdefs - ibody :: "iface \ ibody" +definition ibody :: "iface \ ibody" where "ibody i \ \access=access i,imethods=imethods i\" lemma access_ibody [simp]: "(access (ibody i)) = access i" @@ -351,8 +349,7 @@ super::qtname,superIfs::qtname list,\::'a\" "cdecl" <= (type) "qtname \ class" -constdefs - cbody :: "class \ cbody" +definition cbody :: "class \ cbody" where "cbody c \ \access=access c, cfields=cfields c,methods=methods c,init=init c\" lemma access_cbody [simp]:"access (cbody c) = access c" @@ -394,7 +391,7 @@ lemma SXcptC_inject [simp]: "(SXcptC xn = SXcptC xm) = (xn = xm)" by (simp add: SXcptC_def) -constdefs standard_classes :: "cdecl list" +definition standard_classes :: "cdecl list" where "standard_classes \ [ObjectC, SXcptC Throwable, SXcptC NullPointer, SXcptC OutOfMemory, SXcptC ClassCast, SXcptC NegArrSize , SXcptC IndOutBound, SXcptC ArrStore]" @@ -470,7 +467,7 @@ where "G|-C <:C D == (C,D) \(subcls1 G)^+" notation (xsymbols) - subcls1_syntax ("_\_\\<^sub>C\<^sub>1_" [71,71,71] 70) and + subcls1_syntax ("_\_\\<^sub>C1_" [71,71,71] 70) and subclseq_syntax ("_\_\\<^sub>C _" [71,71,71] 70) and subcls_syntax ("_\_\\<^sub>C _" [71,71,71] 70) @@ -510,7 +507,7 @@ "\G\C \\<^sub>C D\ \ \ c. class G C = Some c" by (auto simp add: subcls1_def dest: tranclD) -lemma no_subcls1_Object:"G\Object\\<^sub>C\<^sub>1 D \ P" +lemma no_subcls1_Object:"G\Object\\<^sub>C1 D \ P" by (auto simp add: subcls1_def) lemma no_subcls_Object: "G\Object\\<^sub>C D \ P" @@ -520,14 +517,13 @@ section "well-structured programs" -constdefs - ws_idecl :: "prog \ qtname \ qtname list \ bool" +definition ws_idecl :: "prog \ qtname \ qtname list \ bool" where "ws_idecl G I si \ \J\set si. is_iface G J \ (J,I)\(subint1 G)^+" - ws_cdecl :: "prog \ qtname \ qtname \ bool" +definition ws_cdecl :: "prog \ qtname \ qtname \ bool" where "ws_cdecl G C sc \ C\Object \ is_class G sc \ (sc,C)\(subcls1 G)^+" - ws_prog :: "prog \ bool" +definition ws_prog :: "prog \ bool" where "ws_prog G \ (\(I,i)\set (ifaces G). ws_idecl G I (isuperIfs i)) \ (\(C,c)\set (classes G). ws_cdecl G C (super c))" @@ -680,7 +676,7 @@ then have "is_class G C \ P C" proof (induct rule: subcls1_induct) fix C - assume hyp:"\ S. G\C \\<^sub>C\<^sub>1 S \ is_class G S \ P S" + assume hyp:"\ S. G\C \\<^sub>C1 S \ is_class G S \ P S" and iscls:"is_class G C" show "P C" proof (cases "C=Object") @@ -715,7 +711,7 @@ then have "\ c. class G C = Some c\ P C c" proof (induct rule: subcls1_induct) fix C c - assume hyp:"\ S. G\C \\<^sub>C\<^sub>1 S \ (\ s. class G S = Some s \ P S s)" + assume hyp:"\ S. G\C \\<^sub>C1 S \ (\ s. class G S = Some s \ P S s)" and iscls:"class G C = Some c" show "P C c" proof (cases "C=Object") @@ -725,7 +721,7 @@ with ws iscls obtain sc where sc: "class G (super c) = Some sc" by (auto dest: ws_prog_cdeclD) - from iscls False have "G\C \\<^sub>C\<^sub>1 (super c)" by (rule subcls1I) + from iscls False have "G\C \\<^sub>C1 (super c)" by (rule subcls1I) with False ws step hyp iscls sc show "P C c" by (auto) @@ -808,8 +804,7 @@ apply simp done -constdefs -imethds:: "prog \ qtname \ (sig,qtname \ mhead) tables" +definition imethds :: "prog \ qtname \ (sig,qtname \ mhead) tables" where --{* methods of an interface, with overriding and inheritance, cf. 9.2 *} "imethds G I \ iface_rec (G,I) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/DeclConcepts.thy --- a/src/HOL/Bali/DeclConcepts.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/DeclConcepts.thy Mon Mar 01 13:42:31 2010 +0100 @@ -8,8 +8,7 @@ section "access control (cf. 6.6), overriding and hiding (cf. 8.4.6.1)" -constdefs -is_public :: "prog \ qtname \ bool" +definition is_public :: "prog \ qtname \ bool" where "is_public G qn \ (case class G qn of None \ (case iface G qn of None \ False @@ -38,14 +37,16 @@ declare accessible_in_RefT_simp [simp del] -constdefs - is_acc_class :: "prog \ pname \ qtname \ bool" +definition is_acc_class :: "prog \ pname \ qtname \ bool" where "is_acc_class G P C \ is_class G C \ G\(Class C) accessible_in P" - is_acc_iface :: "prog \ pname \ qtname \ bool" + +definition is_acc_iface :: "prog \ pname \ qtname \ bool" where "is_acc_iface G P I \ is_iface G I \ G\(Iface I) accessible_in P" - is_acc_type :: "prog \ pname \ ty \ bool" + +definition is_acc_type :: "prog \ pname \ ty \ bool" where "is_acc_type G P T \ is_type G T \ G\T accessible_in P" - is_acc_reftype :: "prog \ pname \ ref_ty \ bool" + +definition is_acc_reftype :: "prog \ pname \ ref_ty \ bool" where "is_acc_reftype G P T \ isrtype G T \ G\T accessible_in' P" lemma is_acc_classD: @@ -336,8 +337,7 @@ text {* Convert a qualified method declaration (qualified with its declaring class) to a qualified member declaration: @{text methdMembr} *} -constdefs -methdMembr :: "(qtname \ mdecl) \ (qtname \ memberdecl)" +definition methdMembr :: "(qtname \ mdecl) \ (qtname \ memberdecl)" where "methdMembr m \ (fst m,mdecl (snd m))" lemma methdMembr_simp[simp]: "methdMembr (c,m) = (c,mdecl m)" @@ -355,8 +355,7 @@ text {* Convert a qualified method (qualified with its declaring class) to a qualified member declaration: @{text method} *} -constdefs -method :: "sig \ (qtname \ methd) \ (qtname \ memberdecl)" +definition method :: "sig \ (qtname \ methd) \ (qtname \ memberdecl)" where "method sig m \ (declclass m, mdecl (sig, mthd m))" lemma method_simp[simp]: "method sig (C,m) = (C,mdecl (sig,m))" @@ -377,8 +376,7 @@ lemma memberid_method_simp[simp]: "memberid (method sig m) = mid sig" by (simp add: method_def) -constdefs -fieldm :: "vname \ (qtname \ field) \ (qtname \ memberdecl)" +definition fieldm :: "vname \ (qtname \ field) \ (qtname \ memberdecl)" where "fieldm n f \ (declclass f, fdecl (n, fld f))" lemma fieldm_simp[simp]: "fieldm n (C,f) = (C,fdecl (n,f))" @@ -402,7 +400,7 @@ text {* Select the signature out of a qualified method declaration: @{text msig} *} -constdefs msig:: "(qtname \ mdecl) \ sig" +definition msig :: "(qtname \ mdecl) \ sig" where "msig m \ fst (snd m)" lemma msig_simp[simp]: "msig (c,(s,m)) = s" @@ -411,7 +409,7 @@ text {* Convert a qualified method (qualified with its declaring class) to a qualified method declaration: @{text qmdecl} *} -constdefs qmdecl :: "sig \ (qtname \ methd) \ (qtname \ mdecl)" +definition qmdecl :: "sig \ (qtname \ methd) \ (qtname \ mdecl)" where "qmdecl sig m \ (declclass m, (sig,mthd m))" lemma qmdecl_simp[simp]: "qmdecl sig (C,m) = (C,(sig,m))" @@ -504,10 +502,8 @@ it is not accessible for inheritance at all. \end{itemize} *} -constdefs -inheritable_in:: - "prog \ (qtname \ memberdecl) \ pname \ bool" - ("_ \ _ inheritable'_in _" [61,61,61] 60) +definition inheritable_in :: "prog \ (qtname \ memberdecl) \ pname \ bool" ("_ \ _ inheritable'_in _" [61,61,61] 60) where + "G\membr inheritable_in pack \ (case (accmodi membr) of Private \ False @@ -529,25 +525,21 @@ subsubsection "declared-in/undeclared-in" -constdefs cdeclaredmethd:: "prog \ qtname \ (sig,methd) table" +definition cdeclaredmethd :: "prog \ qtname \ (sig,methd) table" where "cdeclaredmethd G C \ (case class G C of None \ \ sig. None | Some c \ table_of (methods c) )" -constdefs -cdeclaredfield:: "prog \ qtname \ (vname,field) table" +definition cdeclaredfield :: "prog \ qtname \ (vname,field) table" where "cdeclaredfield G C \ (case class G C of None \ \ sig. None | Some c \ table_of (cfields c) )" - -constdefs -declared_in:: "prog \ memberdecl \ qtname \ bool" - ("_\ _ declared'_in _" [61,61,61] 60) +definition declared_in :: "prog \ memberdecl \ qtname \ bool" ("_\ _ declared'_in _" [61,61,61] 60) where "G\m declared_in C \ (case m of fdecl (fn,f ) \ cdeclaredfield G C fn = Some f | mdecl (sig,m) \ cdeclaredmethd G C sig = Some m)" @@ -567,10 +559,7 @@ by (cases m) (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def) -constdefs -undeclared_in:: "prog \ memberid \ qtname \ bool" - ("_\ _ undeclared'_in _" [61,61,61] 60) - +definition undeclared_in :: "prog \ memberid \ qtname \ bool" ("_\ _ undeclared'_in _" [61,61,61] 60) where "G\m undeclared_in C \ (case m of fid fn \ cdeclaredfield G C fn = None | mid sig \ cdeclaredmethd G C sig = None)" @@ -591,7 +580,7 @@ Immediate: "\G\mbr m declared_in C;declclass m = C\ \ G\m member_of C" | Inherited: "\G\m inheritable_in (pid C); G\memberid m undeclared_in C; - G\C \\<^sub>C\<^sub>1 S; G\(Class S) accessible_in (pid C);G\m member_of S + G\C \\<^sub>C1 S; G\(Class S) accessible_in (pid C);G\m member_of S \ \ G\m member_of C" text {* Note that in the case of an inherited member only the members of the direct superclass are concerned. If a member of a superclass of the direct @@ -617,19 +606,16 @@ ("_ \Field _ _ member'_of _" [61,61,61] 60) where "G\Field n f member_of C == G\fieldm n f member_of C" -constdefs -inherits:: "prog \ qtname \ (qtname \ memberdecl) \ bool" - ("_ \ _ inherits _" [61,61,61] 60) +definition inherits :: "prog \ qtname \ (qtname \ memberdecl) \ bool" ("_ \ _ inherits _" [61,61,61] 60) where "G\C inherits m \ G\m inheritable_in (pid C) \ G\memberid m undeclared_in C \ - (\ S. G\C \\<^sub>C\<^sub>1 S \ G\(Class S) accessible_in (pid C) \ G\m member_of S)" + (\ S. G\C \\<^sub>C1 S \ G\(Class S) accessible_in (pid C) \ G\m member_of S)" lemma inherits_member: "G\C inherits m \ G\m member_of C" by (auto simp add: inherits_def intro: members.Inherited) -constdefs member_in::"prog \ (qtname \ memberdecl) \ qtname \ bool" - ("_ \ _ member'_in _" [61,61,61] 60) +definition member_in :: "prog \ (qtname \ memberdecl) \ qtname \ bool" ("_ \ _ member'_in _" [61,61,61] 60) where "G\m member_in C \ \ provC. G\ C \\<^sub>C provC \ G \ m member_of provC" text {* A member is in a class if it is member of the class or a superclass. If a member is in a class we can select this member. This additional notion @@ -676,7 +662,7 @@ G\Method new declared_in (declclass new); G\Method old declared_in (declclass old); G\Method old inheritable_in pid (declclass new); - G\(declclass new) \\<^sub>C\<^sub>1 superNew; + G\(declclass new) \\<^sub>C1 superNew; G \Method old member_of superNew \ \ G\new overrides\<^sub>S old" @@ -716,9 +702,7 @@ subsubsection "Hiding" -constdefs hides:: -"prog \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" - ("_\ _ hides _" [61,61,61] 60) +definition hides :: "prog \ (qtname \ mdecl) \ (qtname \ mdecl) \ bool" ("_\ _ hides _" [61,61,61] 60) where "G\new hides old \ is_static new \ msig new = msig old \ G\(declclass new) \\<^sub>C (declclass old) \ @@ -777,11 +761,7 @@ by (auto simp add: hides_def) subsubsection "permits access" -constdefs -permits_acc:: - "prog \ (qtname \ memberdecl) \ qtname \ qtname \ bool" - ("_ \ _ in _ permits'_acc'_from _" [61,61,61,61] 60) - +definition permits_acc :: "prog \ (qtname \ memberdecl) \ qtname \ qtname \ bool" ("_ \ _ in _ permits'_acc'_from _" [61,61,61,61] 60) where "G\membr in cls permits_acc_from accclass \ (case (accmodi membr) of Private \ (declclass membr = accclass) @@ -980,7 +960,7 @@ next case (Inherited n C S) assume undecl: "G\ memberid n undeclared_in C" - assume super: "G\C\\<^sub>C\<^sub>1S" + assume super: "G\C\\<^sub>C1S" assume hyp: "\G \ m member_of S; memberid n = memberid m\ \ n = m" assume eqid: "memberid n = memberid m" assume "G \ m member_of C" @@ -1011,7 +991,7 @@ (auto simp add: declared_in_def cdeclaredmethd_def cdeclaredfield_def) next case (Inherited m C S) - assume "G\C\\<^sub>C\<^sub>1S" and "is_class G S" + assume "G\C\\<^sub>C1S" and "is_class G S" then show "is_class G C" by - (rule subcls_is_class2,auto) qed @@ -1043,7 +1023,7 @@ intro: rtrancl_trans) lemma stat_override_declclasses_relation: -"\G\(declclass new) \\<^sub>C\<^sub>1 superNew; G \Method old member_of superNew \ +"\G\(declclass new) \\<^sub>C1 superNew; G \Method old member_of superNew \ \ G\(declclass new) \\<^sub>C (declclass old)" apply (rule trancl_rtrancl_trancl) apply (erule r_into_trancl) @@ -1257,7 +1237,7 @@ "G\ memberid m undeclared_in D" "G \ Class S accessible_in pid D" "G \ m member_of S" - assume super: "G\D\\<^sub>C\<^sub>1S" + assume super: "G\D\\<^sub>C1S" assume hyp: "\G\S\\<^sub>C C; G\C\\<^sub>C declclass m\ \ G \ m member_of C" assume subclseq_C_m: "G\C\\<^sub>C declclass m" assume "G\D\\<^sub>C C" @@ -1399,24 +1379,21 @@ translations "fspec" <= (type) "vname \ qtname" -constdefs -imethds:: "prog \ qtname \ (sig,qtname \ mhead) tables" +definition imethds :: "prog \ qtname \ (sig,qtname \ mhead) tables" where "imethds G I \ iface_rec (G,I) (\I i ts. (Un_tables ts) \\ (Option.set \ table_of (map (\(s,m). (s,I,m)) (imethods i))))" text {* methods of an interface, with overriding and inheritance, cf. 9.2 *} -constdefs -accimethds:: "prog \ pname \ qtname \ (sig,qtname \ mhead) tables" +definition accimethds :: "prog \ pname \ qtname \ (sig,qtname \ mhead) tables" where "accimethds G pack I \ if G\Iface I accessible_in pack then imethds G I else \ k. {}" text {* only returns imethds if the interface is accessible *} -constdefs -methd:: "prog \ qtname \ (sig,qtname \ methd) table" +definition methd :: "prog \ qtname \ (sig,qtname \ methd) table" where "methd G C \ class_rec (G,C) empty @@ -1431,8 +1408,7 @@ Every new method with the same signature coalesces the method of a superclass. *} -constdefs -accmethd:: "prog \ qtname \ qtname \ (sig,qtname \ methd) table" +definition accmethd :: "prog \ qtname \ qtname \ (sig,qtname \ methd) table" where "accmethd G S C \ filter_tab (\sig m. G\method sig m of C accessible_from S) (methd G C)" @@ -1446,8 +1422,7 @@ So we must test accessibility of method @{term m} of class @{term C} (not @{term "declclass m"}) *} -constdefs -dynmethd:: "prog \ qtname \ qtname \ (sig,qtname \ methd) table" +definition dynmethd :: "prog \ qtname \ qtname \ (sig,qtname \ methd) table" where "dynmethd G statC dynC \ \ sig. (if G\dynC \\<^sub>C statC @@ -1473,8 +1448,7 @@ filters the new methods (to get only those methods which actually override the methods of the static class) *} -constdefs -dynimethd:: "prog \ qtname \ qtname \ (sig,qtname \ methd) table" +definition dynimethd :: "prog \ qtname \ qtname \ (sig,qtname \ methd) table" where "dynimethd G I dynC \ \ sig. if imethds G I sig \ {} then methd G dynC sig @@ -1493,8 +1467,7 @@ down to the actual dynamic class. *} -constdefs -dynlookup::"prog \ ref_ty \ qtname \ (sig,qtname \ methd) table" +definition dynlookup :: "prog \ ref_ty \ qtname \ (sig,qtname \ methd) table" where "dynlookup G statT dynC \ (case statT of NullT \ empty @@ -1506,8 +1479,7 @@ In a wellformd context statT will not be NullT and in case statT is an array type, dynC=Object *} -constdefs -fields:: "prog \ qtname \ ((vname \ qtname) \ field) list" +definition fields :: "prog \ qtname \ ((vname \ qtname) \ field) list" where "fields G C \ class_rec (G,C) [] (\C c ts. map (\(n,t). ((n,C),t)) (cfields c) @ ts)" text {* @{term "fields G C"} @@ -1515,8 +1487,7 @@ (private, inherited and hidden ones) not only the accessible ones (an instance of a object allocates all these fields *} -constdefs -accfield:: "prog \ qtname \ qtname \ (vname, qtname \ field) table" +definition accfield :: "prog \ qtname \ qtname \ (vname, qtname \ field) table" where "accfield G S C \ let field_tab = table_of((map (\((n,d),f).(n,(d,f)))) (fields G C)) in filter_tab (\n (declC,f). G\ (declC,fdecl (n,f)) of C accessible_from S) @@ -1531,12 +1502,10 @@ inheritance, too. So we must test accessibility of field @{term f} of class @{term C} (not @{term "declclass f"}) *} -constdefs - - is_methd :: "prog \ qtname \ sig \ bool" +definition is_methd :: "prog \ qtname \ sig \ bool" where "is_methd G \ \C sig. is_class G C \ methd G C sig \ None" -constdefs efname:: "((vname \ qtname) \ field) \ (vname \ qtname)" +definition efname :: "((vname \ qtname) \ field) \ (vname \ qtname)" where "efname \ fst" lemma efname_simp[simp]:"efname (n,f) = n" @@ -2300,8 +2269,7 @@ section "calculation of the superclasses of a class" -constdefs - superclasses:: "prog \ qtname \ qtname set" +definition superclasses :: "prog \ qtname \ qtname set" where "superclasses G C \ class_rec (G,C) {} (\ C c superclss. (if C=Object then {} diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/DefiniteAssignment.thy --- a/src/HOL/Bali/DefiniteAssignment.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/DefiniteAssignment.thy Mon Mar 01 13:42:31 2010 +0100 @@ -74,7 +74,7 @@ "jumpNestingOkS jmps (FinA a c) = False" -constdefs jumpNestingOk :: "jump set \ term \ bool" +definition jumpNestingOk :: "jump set \ term \ bool" where "jumpNestingOk jmps t \ (case t of In1 se \ (case se of Inl e \ True @@ -156,7 +156,7 @@ "assignsEs [] = {}" "assignsEs (e#es) = assignsE e \ assignsEs es" -constdefs assigns:: "term \ lname set" +definition assigns :: "term \ lname set" where "assigns t \ (case t of In1 se \ (case se of Inl e \ assignsE e @@ -429,20 +429,14 @@ subsection {* Lifting set operations to range of tables (map to a set) *} -constdefs - union_ts:: "('a,'b) tables \ ('a,'b) tables \ ('a,'b) tables" - ("_ \\ _" [67,67] 65) +definition union_ts :: "('a,'b) tables \ ('a,'b) tables \ ('a,'b) tables" ("_ \\ _" [67,67] 65) where "A \\ B \ \ k. A k \ B k" -constdefs - intersect_ts:: "('a,'b) tables \ ('a,'b) tables \ ('a,'b) tables" - ("_ \\ _" [72,72] 71) +definition intersect_ts :: "('a,'b) tables \ ('a,'b) tables \ ('a,'b) tables" ("_ \\ _" [72,72] 71) where "A \\ B \ \ k. A k \ B k" -constdefs - all_union_ts:: "('a,'b) tables \ 'b set \ ('a,'b) tables" - (infixl "\\\<^sub>\" 40) -"A \\\<^sub>\ B \ \ k. A k \ B" +definition all_union_ts :: "('a,'b) tables \ 'b set \ ('a,'b) tables" (infixl "\\\<^sub>\" 40) where + "A \\\<^sub>\ B \ \ k. A k \ B" subsubsection {* Binary union of tables *} @@ -513,15 +507,15 @@ brk :: "breakass" --{* Definetly assigned variables for abrupt completion with a break *} -constdefs rmlab :: "'a \ ('a,'b) tables \ ('a,'b) tables" +definition rmlab :: "'a \ ('a,'b) tables \ ('a,'b) tables" where "rmlab k A \ \ x. if x=k then UNIV else A x" (* -constdefs setbrk :: "breakass \ assigned \ breakass set" +definition setbrk :: "breakass \ assigned \ breakass set" where "setbrk b A \ {b} \ {a| a. a\ brk A \ lab a \ lab b}" *) -constdefs range_inter_ts :: "('a,'b) tables \ 'b set" ("\\_" 80) +definition range_inter_ts :: "('a,'b) tables \ 'b set" ("\\_" 80) where "\\A \ {x |x. \ k. x \ A k}" text {* diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/Eval.thy --- a/src/HOL/Bali/Eval.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/Eval.thy Mon Mar 01 13:42:31 2010 +0100 @@ -140,8 +140,7 @@ lst_inj_vals ("\_\\<^sub>l" 1000) where "\es\\<^sub>l == In3 es" -constdefs - undefined3 :: "('al + 'ar, 'b, 'c) sum3 \ vals" +definition undefined3 :: "('al + 'ar, 'b, 'c) sum3 \ vals" where "undefined3 \ sum3_case (In1 \ sum_case (\x. undefined) (\x. Unit)) (\x. In2 undefined) (\x. In3 undefined)" @@ -160,8 +159,7 @@ section "exception throwing and catching" -constdefs - throw :: "val \ abopt \ abopt" +definition throw :: "val \ abopt \ abopt" where "throw a' x \ abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)" lemma throw_def2: @@ -170,8 +168,7 @@ apply (simp (no_asm)) done -constdefs - fits :: "prog \ st \ val \ ty \ bool" ("_,_\_ fits _"[61,61,61,61]60) +definition fits :: "prog \ st \ val \ ty \ bool" ("_,_\_ fits _"[61,61,61,61]60) where "G,s\a' fits T \ (\rt. T=RefT rt) \ a'=Null \ G\obj_ty(lookup_obj s a')\T" lemma fits_Null [simp]: "G,s\Null fits T" @@ -195,8 +192,7 @@ apply iprover done -constdefs - catch ::"prog \ state \ qtname \ bool" ("_,_\catch _"[61,61,61]60) +definition catch :: "prog \ state \ qtname \ bool" ("_,_\catch _"[61,61,61]60) where "G,s\catch C\\xc. abrupt s=Some (Xcpt xc) \ G,store s\Addr (the_Loc xc) fits Class C" @@ -221,8 +217,7 @@ apply (simp (no_asm)) done -constdefs - new_xcpt_var :: "vname \ state \ state" +definition new_xcpt_var :: "vname \ state \ state" where "new_xcpt_var vn \ \(x,s). Norm (lupd(VName vn\Addr (the_Loc (the_Xcpt (the x)))) s)" @@ -237,9 +232,7 @@ section "misc" -constdefs - - assign :: "('a \ state \ state) \ 'a \ state \ state" +definition assign :: "('a \ state \ state) \ 'a \ state \ state" where "assign f v \ \(x,s). let (x',s') = (if x = None then f v else id) (x,s) in (x',if x' = None then s' else s)" @@ -300,9 +293,7 @@ done *) -constdefs - - init_comp_ty :: "ty \ stmt" +definition init_comp_ty :: "ty \ stmt" where "init_comp_ty T \ if (\C. T = Class C) then Init (the_Class T) else Skip" lemma init_comp_ty_PrimT [simp]: "init_comp_ty (PrimT pt) = Skip" @@ -310,9 +301,7 @@ apply (simp (no_asm)) done -constdefs - - invocation_class :: "inv_mode \ st \ val \ ref_ty \ qtname" +definition invocation_class :: "inv_mode \ st \ val \ ref_ty \ qtname" where "invocation_class m s a' statT \ (case m of Static \ if (\ statC. statT = ClassT statC) @@ -321,7 +310,7 @@ | SuperM \ the_Class (RefT statT) | IntVir \ obj_class (lookup_obj s a'))" -invocation_declclass::"prog \ inv_mode \ st \ val \ ref_ty \ sig \ qtname" +definition invocation_declclass::"prog \ inv_mode \ st \ val \ ref_ty \ sig \ qtname" where "invocation_declclass G m s a' statT sig \ declclass (the (dynlookup G statT (invocation_class m s a' statT) @@ -341,9 +330,8 @@ else Object)" by (simp add: invocation_class_def) -constdefs - init_lvars :: "prog \ qtname \ sig \ inv_mode \ val \ val list \ - state \ state" +definition init_lvars :: "prog \ qtname \ sig \ inv_mode \ val \ val list \ + state \ state" where "init_lvars G C sig mode a' pvs \ \ (x,s). let m = mthd (the (methd G C sig)); @@ -376,8 +364,7 @@ apply (simp (no_asm) add: Let_def) done -constdefs - body :: "prog \ qtname \ sig \ expr" +definition body :: "prog \ qtname \ sig \ expr" where "body G C sig \ let m = the (methd G C sig) in Body (declclass m) (stmt (mbody (mthd m)))" @@ -390,12 +377,10 @@ section "variables" -constdefs - - lvar :: "lname \ st \ vvar" +definition lvar :: "lname \ st \ vvar" where "lvar vn s \ (the (locals s vn), \v. supd (lupd(vn\v)))" - fvar :: "qtname \ bool \ vname \ val \ state \ vvar \ state" +definition fvar :: "qtname \ bool \ vname \ val \ state \ vvar \ state" where "fvar C stat fn a' s \ let (oref,xf) = if stat then (Stat C,id) else (Heap (the_Addr a'),np a'); @@ -403,7 +388,7 @@ f = (\v. supd (upd_gobj oref n v)) in ((the (values (the (globs (store s) oref)) n),f),abupd xf s)" - avar :: "prog \ val \ val \ state \ vvar \ state" +definition avar :: "prog \ val \ val \ state \ vvar \ state" where "avar G i' a' s \ let oref = Heap (the_Addr a'); i = the_Intg i'; @@ -446,9 +431,7 @@ apply (simp (no_asm) add: Let_def split_beta) done -constdefs -check_field_access:: -"prog \ qtname \ qtname \ vname \ bool \ val \ state \ state" +definition check_field_access :: "prog \ qtname \ qtname \ vname \ bool \ val \ state \ state" where "check_field_access G accC statDeclC fn stat a' s \ let oref = if stat then Stat statDeclC else Heap (the_Addr a'); @@ -461,9 +444,7 @@ AccessViolation) s" -constdefs -check_method_access:: - "prog \ qtname \ ref_ty \ inv_mode \ sig \ val \ state \ state" +definition check_method_access :: "prog \ qtname \ ref_ty \ inv_mode \ sig \ val \ state \ state" where "check_method_access G accC statT mode sig a' s \ let invC = invocation_class mode (store s) a' statT; dynM = the (dynlookup G statT invC sig) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/Example.thy --- a/src/HOL/Bali/Example.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/Example.thy Mon Mar 01 13:42:31 2010 +0100 @@ -153,23 +153,18 @@ foo :: mname -constdefs - - foo_sig :: sig - "foo_sig \ \name=foo,parTs=[Class Base]\" +definition foo_sig :: sig + where "foo_sig \ \name=foo,parTs=[Class Base]\" - foo_mhead :: mhead - "foo_mhead \ \access=Public,static=False,pars=[z],resT=Class Base\" +definition foo_mhead :: mhead + where "foo_mhead \ \access=Public,static=False,pars=[z],resT=Class Base\" -constdefs - - Base_foo :: mdecl - "Base_foo \ (foo_sig, \access=Public,static=False,pars=[z],resT=Class Base, +definition Base_foo :: mdecl + where "Base_foo \ (foo_sig, \access=Public,static=False,pars=[z],resT=Class Base, mbody=\lcls=[],stmt=Return (!!z)\\)" -constdefs - Ext_foo :: mdecl - "Ext_foo \ (foo_sig, +definition Ext_foo :: mdecl + where "Ext_foo \ (foo_sig, \access=Public,static=False,pars=[z],resT=Class Ext, mbody=\lcls=[] ,stmt=Expr({Ext,Ext,False}Cast (Class Ext) (!!z)..vee := @@ -177,12 +172,10 @@ Return (Lit Null)\ \)" -constdefs - -arr_viewed_from :: "qtname \ qtname \ var" +definition arr_viewed_from :: "qtname \ qtname \ var" where "arr_viewed_from accC C \ {accC,Base,True}StatRef (ClassT C)..arr" -BaseCl :: "class" +definition BaseCl :: "class" where "BaseCl \ \access=Public, cfields=[(arr, \access=Public,static=True ,type=PrimT Boolean.[]\), (vee, \access=Public,static=False,type=Iface HasFoo \)], @@ -192,7 +185,7 @@ super=Object, superIfs=[HasFoo]\" -ExtCl :: "class" +definition ExtCl :: "class" where "ExtCl \ \access=Public, cfields=[(vee, \access=Public,static=False,type= PrimT Integer\)], methods=[Ext_foo], @@ -200,7 +193,7 @@ super=Base, superIfs=[]\" -MainCl :: "class" +definition MainCl :: "class" where "MainCl \ \access=Public, cfields=[], methods=[], @@ -209,16 +202,14 @@ superIfs=[]\" (* The "main" method is modeled seperately (see tprg) *) -constdefs - - HasFooInt :: iface - "HasFooInt \ \access=Public,imethods=[(foo_sig, foo_mhead)],isuperIfs=[]\" +definition HasFooInt :: iface + where "HasFooInt \ \access=Public,imethods=[(foo_sig, foo_mhead)],isuperIfs=[]\" - Ifaces ::"idecl list" - "Ifaces \ [(HasFoo,HasFooInt)]" +definition Ifaces ::"idecl list" + where "Ifaces \ [(HasFoo,HasFooInt)]" - "Classes" ::"cdecl list" - "Classes \ [(Base,BaseCl),(Ext,ExtCl),(Main,MainCl)]@standard_classes" +definition "Classes" ::"cdecl list" + where "Classes \ [(Base,BaseCl),(Ext,ExtCl),(Main,MainCl)]@standard_classes" lemmas table_classes_defs = Classes_def standard_classes_def ObjectC_def SXcptC_def @@ -273,8 +264,7 @@ tprg :: prog where "tprg == \ifaces=Ifaces,classes=Classes\" -constdefs - test :: "(ty)list \ stmt" +definition test :: "(ty)list \ stmt" where "test pTs \ e:==NewC Ext;; \ Try Expr({Main,ClassT Base,IntVir}!!e\foo({pTs}[Lit Null])) \ Catch((SXcpt NullPointer) z) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/State.thy --- a/src/HOL/Bali/State.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/State.thy Mon Mar 01 13:42:31 2010 +0100 @@ -38,9 +38,7 @@ "obj" <= (type) "\tag::obj_tag, values::vn \ val option\" "obj" <= (type) "\tag::obj_tag, values::vn \ val option,\::'a\" -constdefs - - the_Arr :: "obj option \ ty \ int \ (vn, val) table" +definition the_Arr :: "obj option \ ty \ int \ (vn, val) table" where "the_Arr obj \ SOME (T,k,t). obj = Some \tag=Arr T k,values=t\" lemma the_Arr_Arr [simp]: "the_Arr (Some \tag=Arr T k,values=cs\) = (T,k,cs)" @@ -52,9 +50,7 @@ apply (auto simp add: the_Arr_def) done -constdefs - - upd_obj :: "vn \ val \ obj \ obj" +definition upd_obj :: "vn \ val \ obj \ obj" where "upd_obj n v \ \ obj . obj \values:=(values obj)(n\v)\" lemma upd_obj_def2 [simp]: @@ -62,8 +58,7 @@ apply (auto simp: upd_obj_def) done -constdefs - obj_ty :: "obj \ ty" +definition obj_ty :: "obj \ ty" where "obj_ty obj \ case tag obj of CInst C \ Class C | Arr T k \ T.[]" @@ -102,9 +97,7 @@ apply (auto split add: obj_tag.split_asm) done -constdefs - - obj_class :: "obj \ qtname" +definition obj_class :: "obj \ qtname" where "obj_class obj \ case tag obj of CInst C \ C | Arr T k \ Object" @@ -143,9 +136,7 @@ "Stat" => "CONST Inr" "oref" <= (type) "loc + qtname" -constdefs - fields_table:: - "prog \ qtname \ (fspec \ field \ bool) \ (fspec, ty) table" +definition fields_table :: "prog \ qtname \ (fspec \ field \ bool) \ (fspec, ty) table" where "fields_table G C P \ Option.map type \ table_of (filter (split P) (DeclConcepts.fields G C))" @@ -182,14 +173,13 @@ apply simp done -constdefs - in_bounds :: "int \ int \ bool" ("(_/ in'_bounds _)" [50, 51] 50) +definition in_bounds :: "int \ int \ bool" ("(_/ in'_bounds _)" [50, 51] 50) where "i in_bounds k \ 0 \ i \ i < k" - arr_comps :: "'a \ int \ int \ 'a option" +definition arr_comps :: "'a \ int \ int \ 'a option" where "arr_comps T k \ \i. if i in_bounds k then Some T else None" - var_tys :: "prog \ obj_tag \ oref \ (vn, ty) table" +definition var_tys :: "prog \ obj_tag \ oref \ (vn, ty) table" where "var_tys G oi r \ case r of Heap a \ (case oi of @@ -232,15 +222,13 @@ subsection "access" -constdefs - - globs :: "st \ globs" +definition globs :: "st \ globs" where "globs \ st_case (\g l. g)" - locals :: "st \ locals" +definition locals :: "st \ locals" where "locals \ st_case (\g l. l)" - heap :: "st \ heap" +definition heap :: "st \ heap" where "heap s \ globs s \ Heap" @@ -262,8 +250,7 @@ subsection "memory allocation" -constdefs - new_Addr :: "heap \ loc option" +definition new_Addr :: "heap \ loc option" where "new_Addr h \ if (\a. h a \ None) then None else Some (SOME a. h a = None)" lemma new_AddrD: "new_Addr h = Some a \ h a = None" @@ -303,20 +290,19 @@ subsection "update" -constdefs - gupd :: "oref \ obj \ st \ st" ("gupd'(_\_')"[10,10]1000) +definition gupd :: "oref \ obj \ st \ st" ("gupd'(_\_')"[10,10]1000) where "gupd r obj \ st_case (\g l. st (g(r\obj)) l)" - lupd :: "lname \ val \ st \ st" ("lupd'(_\_')"[10,10]1000) +definition lupd :: "lname \ val \ st \ st" ("lupd'(_\_')"[10,10]1000) where "lupd vn v \ st_case (\g l. st g (l(vn\v)))" - upd_gobj :: "oref \ vn \ val \ st \ st" +definition upd_gobj :: "oref \ vn \ val \ st \ st" where "upd_gobj r n v \ st_case (\g l. st (chg_map (upd_obj n v) r g) l)" - set_locals :: "locals \ st \ st" +definition set_locals :: "locals \ st \ st" where "set_locals l \ st_case (\g l'. st g l)" - init_obj :: "prog \ obj_tag \ oref \ st \ st" +definition init_obj :: "prog \ obj_tag \ oref \ st \ st" where "init_obj G oi r \ gupd(r\\tag=oi, values=init_vals (var_tys G oi r)\)" abbreviation @@ -476,8 +462,7 @@ -constdefs - abrupt_if :: "bool \ abopt \ abopt \ abopt" +definition abrupt_if :: "bool \ abopt \ abopt \ abopt" where "abrupt_if c x' x \ if c \ (x = None) then x' else x" lemma abrupt_if_True_None [simp]: "abrupt_if True x None = x" @@ -557,8 +542,7 @@ apply auto done -constdefs - absorb :: "jump \ abopt \ abopt" +definition absorb :: "jump \ abopt \ abopt" where "absorb j a \ if a=Some (Jump j) then None else a" lemma absorb_SomeD [dest!]: "absorb j a = Some x \ a = Some x" @@ -611,9 +595,7 @@ apply clarsimp done -constdefs - - normal :: "state \ bool" +definition normal :: "state \ bool" where "normal \ \s. abrupt s = None" lemma normal_def2 [simp]: "normal s = (abrupt s = None)" @@ -621,8 +603,7 @@ apply (simp (no_asm)) done -constdefs - heap_free :: "nat \ state \ bool" +definition heap_free :: "nat \ state \ bool" where "heap_free n \ \s. atleast_free (heap (store s)) n" lemma heap_free_def2 [simp]: "heap_free n s = atleast_free (heap (store s)) n" @@ -632,12 +613,10 @@ subsection "update" -constdefs - - abupd :: "(abopt \ abopt) \ state \ state" +definition abupd :: "(abopt \ abopt) \ state \ state" where "abupd f \ prod_fun f id" - supd :: "(st \ st) \ state \ state" +definition supd :: "(st \ st) \ state \ state" where "supd \ prod_fun id" lemma abupd_def2 [simp]: "abupd f (x,s) = (f x,s)" @@ -692,12 +671,10 @@ section "initialisation test" -constdefs - - inited :: "qtname \ globs \ bool" +definition inited :: "qtname \ globs \ bool" where "inited C g \ g (Stat C) \ None" - initd :: "qtname \ state \ bool" +definition initd :: "qtname \ state \ bool" where "initd C \ inited C \ globs \ store" lemma not_inited_empty [simp]: "\inited C empty" @@ -731,7 +708,7 @@ done section {* @{text error_free} *} -constdefs error_free:: "state \ bool" +definition error_free :: "state \ bool" where "error_free s \ \ (\ err. abrupt s = Some (Error err))" lemma error_free_Norm [simp,intro]: "error_free (Norm s)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/Table.thy --- a/src/HOL/Bali/Table.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/Table.thy Mon Mar 01 13:42:31 2010 +0100 @@ -51,9 +51,7 @@ by (simp add: map_add_def) section {* Conditional Override *} -constdefs -cond_override:: - "('b \'b \ bool) \ ('a, 'b)table \ ('a, 'b)table \ ('a, 'b) table" +definition cond_override :: "('b \'b \ bool) \ ('a, 'b)table \ ('a, 'b)table \ ('a, 'b) table" where --{* when merging tables old and new, only override an entry of table old when the condition cond holds *} @@ -98,8 +96,7 @@ section {* Filter on Tables *} -constdefs -filter_tab:: "('a \ 'b \ bool) \ ('a, 'b) table \ ('a, 'b) table" +definition filter_tab :: "('a \ 'b \ bool) \ ('a, 'b) table \ ('a, 'b) table" where "filter_tab c t \ \ k. (case t k of None \ None | Some x \ if c k x then Some x else None)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/Term.thy --- a/src/HOL/Bali/Term.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/Term.thy Mon Mar 01 13:42:31 2010 +0100 @@ -261,9 +261,7 @@ StatRef :: "ref_ty \ expr" where "StatRef rt == Cast (RefT rt) (Lit Null)" -constdefs - - is_stmt :: "term \ bool" +definition is_stmt :: "term \ bool" where "is_stmt t \ \c. t=In1r c" ML {* bind_thms ("is_stmt_rews", sum3_instantiate @{context} @{thm is_stmt_def}) *} @@ -467,7 +465,7 @@ "eval_binop CondAnd v1 v2 = Bool ((the_Bool v1) \ (the_Bool v2))" "eval_binop CondOr v1 v2 = Bool ((the_Bool v1) \ (the_Bool v2))" -constdefs need_second_arg :: "binop \ val \ bool" +definition need_second_arg :: "binop \ val \ bool" where "need_second_arg binop v1 \ \ ((binop=CondAnd \ \ the_Bool v1) \ (binop=CondOr \ the_Bool v1))" text {* @{term CondAnd} and @{term CondOr} only evalulate the second argument diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/Trans.thy --- a/src/HOL/Bali/Trans.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/Trans.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,7 +9,7 @@ theory Trans imports Evaln begin -constdefs groundVar:: "var \ bool" +definition groundVar :: "var \ bool" where "groundVar v \ (case v of LVar ln \ True | {accC,statDeclC,stat}e..fn \ \ a. e=Lit a @@ -34,7 +34,7 @@ done qed -constdefs groundExprs:: "expr list \ bool" +definition groundExprs :: "expr list \ bool" where "groundExprs es \ list_all (\ e. \ v. e=Lit v) es" consts the_val:: "expr \ val" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/Type.thy --- a/src/HOL/Bali/Type.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/Type.thy Mon Mar 01 13:42:31 2010 +0100 @@ -41,8 +41,7 @@ abbreviation Array :: "ty \ ty" ("_.[]" [90] 90) where "T.[] == RefT (ArrayT T)" -constdefs - the_Class :: "ty \ qtname" +definition the_Class :: "ty \ qtname" where "the_Class T \ SOME C. T = Class C" (** primrec not possible here **) lemma the_Class_eq [simp]: "the_Class (Class C)= C" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/TypeRel.thy --- a/src/HOL/Bali/TypeRel.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/TypeRel.thy Mon Mar 01 13:42:31 2010 +0100 @@ -65,7 +65,7 @@ done lemma subcls1I1: - "\class G C = Some c; C \ Object;D=super c\ \ G\C\\<^sub>C\<^sub>1 D" + "\class G C = Some c; C \ Object;D=super c\ \ G\C\\<^sub>C1 D" apply (auto dest: subcls1I) done @@ -126,7 +126,7 @@ done lemma single_inheritance: -"\G\A \\<^sub>C\<^sub>1 B; G\A \\<^sub>C\<^sub>1 C\ \ B = C" +"\G\A \\<^sub>C1 B; G\A \\<^sub>C1 C\ \ B = C" by (auto simp add: subcls1_def) lemma subcls_compareable: @@ -134,11 +134,11 @@ \ \ G\X \\<^sub>C Y \ G\Y \\<^sub>C X" by (rule triangle_lemma) (auto intro: single_inheritance) -lemma subcls1_irrefl: "\G\C \\<^sub>C\<^sub>1 D; ws_prog G \ +lemma subcls1_irrefl: "\G\C \\<^sub>C1 D; ws_prog G \ \ C \ D" proof assume ws: "ws_prog G" and - subcls1: "G\C \\<^sub>C\<^sub>1 D" and + subcls1: "G\C \\<^sub>C1 D" and eq_C_D: "C=D" from subcls1 obtain c where @@ -167,7 +167,7 @@ then show ?thesis proof (induct rule: converse_trancl_induct) fix C - assume subcls1_C_D: "G\C \\<^sub>C\<^sub>1 D" + assume subcls1_C_D: "G\C \\<^sub>C1 D" then obtain c where "C\Object" and "class G C = Some c" and @@ -178,7 +178,7 @@ by (auto dest: ws_prog_cdeclD) next fix C Z - assume subcls1_C_Z: "G\C \\<^sub>C\<^sub>1 Z" and + assume subcls1_C_Z: "G\C \\<^sub>C1 Z" and subcls_Z_D: "G\Z \\<^sub>C D" and nsubcls_D_Z: "\ G\D \\<^sub>C Z" show "\ G\D \\<^sub>C C" @@ -213,13 +213,13 @@ then show ?thesis proof (induct rule: converse_trancl_induct) fix C - assume "G\C \\<^sub>C\<^sub>1 D" + assume "G\C \\<^sub>C1 D" with ws show "C\D" by (blast dest: subcls1_irrefl) next fix C Z - assume subcls1_C_Z: "G\C \\<^sub>C\<^sub>1 Z" and + assume subcls1_C_Z: "G\C \\<^sub>C1 Z" and subcls_Z_D: "G\Z \\<^sub>C D" and neq_Z_D: "Z \ D" show "C\D" @@ -298,7 +298,7 @@ assume clsC: "class G C = Some c" assume subcls_C_C: "G\C \\<^sub>C D" then obtain S where - "G\C \\<^sub>C\<^sub>1 S" and + "G\C \\<^sub>C1 S" and subclseq_S_D: "G\S \\<^sub>C D" by (blast dest: tranclD) with clsC @@ -341,9 +341,9 @@ where direct: "G\C\1J \\ \ G\C\J" | subint: "\G\C\1I; G\I\I J\ \ G\C\J" -| subcls1: "\G\C\\<^sub>C\<^sub>1D; G\D\J \ \ G\C\J" +| subcls1: "\G\C\\<^sub>C1D; G\D\J \ \ G\C\J" -lemma implmtD: "G\C\J \ (\I. G\C\1I \ G\I\I J) \ (\D. G\C\\<^sub>C\<^sub>1D \ G\D\J)" +lemma implmtD: "G\C\J \ (\I. G\C\1I \ G\I\I J) \ (\D. G\C\\<^sub>C1D \ G\D\J)" apply (erule implmt.induct) apply fast+ done @@ -568,8 +568,7 @@ apply (fast dest: widen_Class_Class widen_Class_Iface) done -constdefs - widens :: "prog \ [ty list, ty list] \ bool" ("_\_[\]_" [71,71,71] 70) +definition widens :: "prog \ [ty list, ty list] \ bool" ("_\_[\]_" [71,71,71] 70) where "G\Ts[\]Ts' \ list_all2 (\T T'. G\T\T') Ts Ts'" lemma widens_Nil [simp]: "G\[][\][]" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/TypeSafe.thy --- a/src/HOL/Bali/TypeSafe.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/TypeSafe.thy Mon Mar 01 13:42:31 2010 +0100 @@ -95,17 +95,13 @@ section "result conformance" -constdefs - assign_conforms :: "st \ (val \ state \ state) \ ty \ env' \ bool" - ("_\|_\_\\_" [71,71,71,71] 70) +definition assign_conforms :: "st \ (val \ state \ state) \ ty \ env' \ bool" ("_\|_\_\\_" [71,71,71,71] 70) where "s\|f\T\\E \ (\s' w. Norm s'\\E \ fst E,s'\w\\T \ s\|s' \ assign f w (Norm s')\\E) \ (\s' w. error_free s' \ (error_free (assign f w s')))" -constdefs - rconf :: "prog \ lenv \ st \ term \ vals \ tys \ bool" - ("_,_,_\_\_\\_" [71,71,71,71,71,71] 70) +definition rconf :: "prog \ lenv \ st \ term \ vals \ tys \ bool" ("_,_,_\_\_\\_" [71,71,71,71,71,71] 70) where "G,L,s\t\v\\T \ case T of Inl T \ if (\ var. t=In2 var) @@ -330,11 +326,8 @@ declare fun_upd_apply [simp del] - -constdefs - DynT_prop::"[prog,inv_mode,qtname,ref_ty] \ bool" - ("_\_\_\_"[71,71,71,71]70) - "G\mode\D\t \ mode = IntVir \ is_class G D \ +definition DynT_prop :: "[prog,inv_mode,qtname,ref_ty] \ bool" ("_\_\_\_"[71,71,71,71]70) where + "G\mode\D\t \ mode = IntVir \ is_class G D \ (if (\T. t=ArrayT T) then D=Object else G\Class D\RefT t)" lemma DynT_propI: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/WellForm.thy --- a/src/HOL/Bali/WellForm.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/WellForm.thy Mon Mar 01 13:42:31 2010 +0100 @@ -31,8 +31,7 @@ text {* well-formed field declaration (common part for classes and interfaces), cf. 8.3 and (9.3) *} -constdefs - wf_fdecl :: "prog \ pname \ fdecl \ bool" +definition wf_fdecl :: "prog \ pname \ fdecl \ bool" where "wf_fdecl G P \ \(fn,f). is_acc_type G P (type f)" lemma wf_fdecl_def2: "\fd. wf_fdecl G P fd = is_acc_type G P (type (snd fd))" @@ -55,8 +54,7 @@ \item the parameter names are unique \end{itemize} *} -constdefs - wf_mhead :: "prog \ pname \ sig \ mhead \ bool" +definition wf_mhead :: "prog \ pname \ sig \ mhead \ bool" where "wf_mhead G P \ \ sig mh. length (parTs sig) = length (pars mh) \ \ ( \T\set (parTs sig). is_acc_type G P T) \ is_acc_type G P (resTy mh) \ @@ -78,7 +76,7 @@ \end{itemize} *} -constdefs callee_lcl:: "qtname \ sig \ methd \ lenv" +definition callee_lcl :: "qtname \ sig \ methd \ lenv" where "callee_lcl C sig m \ \ k. (case k of EName e @@ -88,12 +86,11 @@ | Res \ Some (resTy m)) | This \ if is_static m then None else Some (Class C))" -constdefs parameters :: "methd \ lname set" +definition parameters :: "methd \ lname set" where "parameters m \ set (map (EName \ VNam) (pars m)) \ (if (static m) then {} else {This})" -constdefs - wf_mdecl :: "prog \ qtname \ mdecl \ bool" +definition wf_mdecl :: "prog \ qtname \ mdecl \ bool" where "wf_mdecl G C \ \(sig,m). wf_mhead G (pid C) sig (mhead m) \ @@ -219,8 +216,7 @@ superinterfaces widens to each of the corresponding result types \end{itemize} *} -constdefs - wf_idecl :: "prog \ idecl \ bool" +definition wf_idecl :: "prog \ idecl \ bool" where "wf_idecl G \ \(I,i). ws_idecl G I (isuperIfs i) \ @@ -321,8 +317,7 @@ \end{itemize} *} (* to Table *) -constdefs entails:: "('a,'b) table \ ('b \ bool) \ bool" - ("_ entails _" 20) +definition entails :: "('a,'b) table \ ('b \ bool) \ bool" ("_ entails _" 20) where "t entails P \ \k. \ x \ t k: P x" lemma entailsD: @@ -332,8 +327,7 @@ lemma empty_entails[simp]: "empty entails P" by (simp add: entails_def) -constdefs - wf_cdecl :: "prog \ cdecl \ bool" +definition wf_cdecl :: "prog \ cdecl \ bool" where "wf_cdecl G \ \(C,c). \is_iface G C \ @@ -361,8 +355,7 @@ ))" (* -constdefs - wf_cdecl :: "prog \ cdecl \ bool" +definition wf_cdecl :: "prog \ cdecl \ bool" where "wf_cdecl G \ \(C,c). \is_iface G C \ @@ -518,8 +511,7 @@ \item all defined classes are wellformed \end{itemize} *} -constdefs - wf_prog :: "prog \ bool" +definition wf_prog :: "prog \ bool" where "wf_prog G \ let is = ifaces G; cs = classes G in ObjectC \ set cs \ (\ m\set Object_mdecls. accmodi m \ Package) \ @@ -919,7 +911,7 @@ inheritable: "G \Method old inheritable_in pid C" and subclsC: "G\C\\<^sub>C declclass old" from cls_C neq_C_Obj - have super: "G\C \\<^sub>C\<^sub>1 super c" + have super: "G\C \\<^sub>C1 super c" by (rule subcls1I) from wf cls_C neq_C_Obj have accessible_super: "G\(Class (super c)) accessible_in (pid C)" @@ -1385,7 +1377,7 @@ moreover note wf False cls_C ultimately have "G\super c \\<^sub>C declclass m" by (auto intro: Hyp [rule_format]) - moreover from cls_C False have "G\C \\<^sub>C\<^sub>1 super c" by (rule subcls1I) + moreover from cls_C False have "G\C \\<^sub>C1 super c" by (rule subcls1I) ultimately show ?thesis by - (rule rtrancl_into_rtrancl2) next case Some @@ -1539,7 +1531,7 @@ by (auto intro: method_declared_inI) note trancl_rtrancl_tranc = trancl_rtrancl_trancl [trans] (* ### in Basis *) from clsC neq_C_Obj - have subcls1_C_super: "G\C \\<^sub>C\<^sub>1 super c" + have subcls1_C_super: "G\C \\<^sub>C1 super c" by (rule subcls1I) then have "G\C \\<^sub>C super c" .. also from old wf is_cls_super @@ -1609,7 +1601,7 @@ by (auto dest: ws_prog_cdeclD) from clsC wf neq_C_Obj have superAccessible: "G\(Class (super c)) accessible_in (pid C)" and - subcls1_C_super: "G\C \\<^sub>C\<^sub>1 super c" + subcls1_C_super: "G\C \\<^sub>C1 super c" by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD intro: subcls1I) show "\new. ?Constraint C new old" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Bali/WellType.thy --- a/src/HOL/Bali/WellType.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Bali/WellType.thy Mon Mar 01 13:42:31 2010 +0100 @@ -53,11 +53,10 @@ emhead = "ref_ty \ mhead" --{* Some mnemotic selectors for emhead *} -constdefs - "declrefT" :: "emhead \ ref_ty" +definition "declrefT" :: "emhead \ ref_ty" where "declrefT \ fst" - "mhd" :: "emhead \ mhead" +definition "mhd" :: "emhead \ mhead" where "mhd \ snd" lemma declrefT_simp[simp]:"declrefT (r,m) = r" @@ -138,11 +137,10 @@ done -constdefs - empty_dt :: "dyn_ty" +definition empty_dt :: "dyn_ty" where "empty_dt \ \a. None" - invmode :: "('a::type)member_scheme \ expr \ inv_mode" +definition invmode :: "('a::type)member_scheme \ expr \ inv_mode" where "invmode m e \ if is_static m then Static else if e=Super then SuperM else IntVir" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Decision_Procs/Cooper.thy --- a/src/HOL/Decision_Procs/Cooper.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Decision_Procs/Cooper.thy Mon Mar 01 13:42:31 2010 +0100 @@ -293,10 +293,10 @@ by (induct p, simp_all) -constdefs djf:: "('a \ fm) \ 'a \ fm \ fm" +definition djf :: "('a \ fm) \ 'a \ fm \ fm" where "djf f p q \ (if q=T then T else if q=F then f p else (let fp = f p in case fp of T \ T | F \ q | _ \ Or (f p) q))" -constdefs evaldjf:: "('a \ fm) \ 'a list \ fm" +definition evaldjf :: "('a \ fm) \ 'a list \ fm" where "evaldjf f ps \ foldr (djf f) ps F" lemma djf_Or: "Ifm bbs bs (djf f p q) = Ifm bbs bs (Or (f p) q)" @@ -340,7 +340,7 @@ thus ?thesis by (simp only: list_all_iff) qed -constdefs DJ :: "(fm \ fm) \ fm \ fm" +definition DJ :: "(fm \ fm) \ fm \ fm" where "DJ f p \ evaldjf f (disjuncts p)" lemma DJ: assumes fdj: "\ p q. f (Or p q) = Or (f p) (f q)" @@ -395,7 +395,7 @@ "lex_ns ([], ms) = True" "lex_ns (ns, []) = False" "lex_ns (n#ns, m#ms) = (n ((n = m) \ lex_ns (ns,ms))) " -constdefs lex_bnd :: "num \ num \ bool" +definition lex_bnd :: "num \ num \ bool" where "lex_bnd t s \ lex_ns (bnds t, bnds s)" consts @@ -455,10 +455,10 @@ lemma nummul_nb: "\ i. numbound0 t \ numbound0 (nummul i t)" by (induct t rule: nummul.induct, auto simp add: numadd_nb) -constdefs numneg :: "num \ num" +definition numneg :: "num \ num" where "numneg t \ nummul (- 1) t" -constdefs numsub :: "num \ num \ num" +definition numsub :: "num \ num \ num" where "numsub s t \ (if s = t then C 0 else numadd (s, numneg t))" lemma numneg: "Inum bs (numneg t) = Inum bs (Neg t)" @@ -505,7 +505,7 @@ lemma not_bn: "bound0 p \ bound0 (not p)" by (cases p, auto) -constdefs conj :: "fm \ fm \ fm" +definition conj :: "fm \ fm \ fm" where "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else And p q)" lemma conj: "Ifm bbs bs (conj p q) = Ifm bbs bs (And p q)" by (cases "p=F \ q=F",simp_all add: conj_def) (cases p,simp_all) @@ -515,7 +515,7 @@ lemma conj_nb: "\bound0 p ; bound0 q\ \ bound0 (conj p q)" using conj_def by auto -constdefs disj :: "fm \ fm \ fm" +definition disj :: "fm \ fm \ fm" where "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p else Or p q)" lemma disj: "Ifm bbs bs (disj p q) = Ifm bbs bs (Or p q)" @@ -525,7 +525,7 @@ lemma disj_nb: "\bound0 p ; bound0 q\ \ bound0 (disj p q)" using disj_def by auto -constdefs imp :: "fm \ fm \ fm" +definition imp :: "fm \ fm \ fm" where "imp p q \ (if (p = F \ q=T) then T else if p=T then q else if q=F then not p else Imp p q)" lemma imp: "Ifm bbs bs (imp p q) = Ifm bbs bs (Imp p q)" by (cases "p=F \ q=T",simp_all add: imp_def,cases p) (simp_all add: not) @@ -534,7 +534,7 @@ lemma imp_nb: "\bound0 p ; bound0 q\ \ bound0 (imp p q)" using imp_def by (cases "p=F \ q=T",simp_all add: imp_def,cases p) simp_all -constdefs iff :: "fm \ fm \ fm" +definition iff :: "fm \ fm \ fm" where "iff p q \ (if (p = q) then T else if (p = not q \ not p = q) then F else if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else Iff p q)" @@ -1749,7 +1749,7 @@ shows "(\ x. Ifm bbs (x#bs) p) = ((\ j\ {1 .. d}. Ifm bbs (j#bs) (minusinf p)) \ (\ j\ {1.. d}. \ b\ (Inum (i#bs)) ` set (\ p). Ifm bbs ((b+j)#bs) p))" using cp_thm[OF lp up dd dp,where i="i"] by auto -constdefs unit:: "fm \ fm \ num list \ int" +definition unit :: "fm \ fm \ num list \ int" where "unit p \ (let p' = zlfm p ; l = \ p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\ p' l); d = \ q; B = remdups (map simpnum (\ q)) ; a = remdups (map simpnum (\ q)) in if length B \ length a then (q,B,d) else (mirror q, a,d))" @@ -1814,7 +1814,7 @@ qed (* Cooper's Algorithm *) -constdefs cooper :: "fm \ fm" +definition cooper :: "fm \ fm" where "cooper p \ (let (q,B,d) = unit p; js = iupt 1 d; mq = simpfm (minusinf q); diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Decision_Procs/Ferrack.thy --- a/src/HOL/Decision_Procs/Ferrack.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Decision_Procs/Ferrack.thy Mon Mar 01 13:42:31 2010 +0100 @@ -169,26 +169,26 @@ lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)" by (cases p) auto -constdefs conj :: "fm \ fm \ fm" +definition conj :: "fm \ fm \ fm" where "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else if p = q then p else And p q)" lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)" by (cases "p=F \ q=F",simp_all add: conj_def) (cases p,simp_all) -constdefs disj :: "fm \ fm \ fm" +definition disj :: "fm \ fm \ fm" where "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p else if p=q then p else Or p q)" lemma disj[simp]: "Ifm bs (disj p q) = Ifm bs (Or p q)" by (cases "p=T \ q=T",simp_all add: disj_def) (cases p,simp_all) -constdefs imp :: "fm \ fm \ fm" +definition imp :: "fm \ fm \ fm" where "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p else Imp p q)" lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)" by (cases "p=F \ q=T",simp_all add: imp_def) -constdefs iff :: "fm \ fm \ fm" +definition iff :: "fm \ fm \ fm" where "iff p q \ (if (p = q) then T else if (p = NOT q \ NOT p = q) then F else if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else Iff p q)" @@ -369,10 +369,10 @@ lemma bound0_qf: "bound0 p \ qfree p" by (induct p, simp_all) -constdefs djf:: "('a \ fm) \ 'a \ fm \ fm" +definition djf :: "('a \ fm) \ 'a \ fm \ fm" where "djf f p q \ (if q=T then T else if q=F then f p else (let fp = f p in case fp of T \ T | F \ q | _ \ Or (f p) q))" -constdefs evaldjf:: "('a \ fm) \ 'a list \ fm" +definition evaldjf :: "('a \ fm) \ 'a list \ fm" where "evaldjf f ps \ foldr (djf f) ps F" lemma djf_Or: "Ifm bs (djf f p q) = Ifm bs (Or (f p) q)" @@ -423,7 +423,7 @@ thus ?thesis by (simp only: list_all_iff) qed -constdefs DJ :: "(fm \ fm) \ fm \ fm" +definition DJ :: "(fm \ fm) \ fm \ fm" where "DJ f p \ evaldjf f (disjuncts p)" lemma DJ: assumes fdj: "\ p q. Ifm bs (f (Or p q)) = Ifm bs (Or (f p) (f q))" @@ -653,10 +653,10 @@ lemma nummul_nb[simp]: "\ i. numbound0 t \ numbound0 (nummul t i)" by (induct t rule: nummul.induct, auto ) -constdefs numneg :: "num \ num" +definition numneg :: "num \ num" where "numneg t \ nummul t (- 1)" -constdefs numsub :: "num \ num \ num" +definition numsub :: "num \ num \ num" where "numsub s t \ (if s = t then C 0 else numadd (s,numneg t))" lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)" @@ -724,7 +724,7 @@ from maxcoeff_nz[OF nz th] show ?thesis . qed -constdefs simp_num_pair:: "(num \ int) \ num \ int" +definition simp_num_pair :: "(num \ int) \ num \ int" where "simp_num_pair \ (\ (t,n). (if n = 0 then (C 0, 0) else (let t' = simpnum t ; g = numgcd t' in if g > 1 then (let g' = gcd n g in @@ -1779,7 +1779,7 @@ (* Implement the right hand side of Ferrante and Rackoff's Theorem. *) -constdefs ferrack:: "fm \ fm" +definition ferrack :: "fm \ fm" where "ferrack p \ (let p' = rlfm (simpfm p); mp = minusinf p'; pp = plusinf p' in if (mp = T \ pp = T) then T else (let U = remdps(map simp_num_pair diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Decision_Procs/MIR.thy --- a/src/HOL/Decision_Procs/MIR.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Decision_Procs/MIR.thy Mon Mar 01 13:42:31 2010 +0100 @@ -566,7 +566,7 @@ thus ?thesis by (simp only: list_all_iff) qed -constdefs DJ :: "(fm \ fm) \ fm \ fm" +definition DJ :: "(fm \ fm) \ fm \ fm" where "DJ f p \ evaldjf f (disjuncts p)" lemma DJ: assumes fdj: "\ p q. f (Or p q) = Or (f p) (f q)" @@ -623,7 +623,7 @@ "lex_ns ([], ms) = True" "lex_ns (ns, []) = False" "lex_ns (n#ns, m#ms) = (n ((n = m) \ lex_ns (ns,ms))) " -constdefs lex_bnd :: "num \ num \ bool" +definition lex_bnd :: "num \ num \ bool" where "lex_bnd t s \ lex_ns (bnds t, bnds s)" consts @@ -873,10 +873,10 @@ lemma nummul_nb[simp]: "\ i. numbound0 t \ numbound0 (nummul t i)" by (induct t rule: nummul.induct, auto) -constdefs numneg :: "num \ num" +definition numneg :: "num \ num" where "numneg t \ nummul t (- 1)" -constdefs numsub :: "num \ num \ num" +definition numsub :: "num \ num \ num" where "numsub s t \ (if s = t then C 0 else numadd (s,numneg t))" lemma numneg[simp]: "Inum bs (numneg t) = Inum bs (Neg t)" @@ -1038,7 +1038,7 @@ from maxcoeff_nz[OF nz th] show ?thesis . qed -constdefs simp_num_pair:: "(num \ int) \ num \ int" +definition simp_num_pair :: "(num \ int) \ num \ int" where "simp_num_pair \ (\ (t,n). (if n = 0 then (C 0, 0) else (let t' = simpnum t ; g = numgcd t' in if g > 1 then (let g' = gcd n g in @@ -1137,7 +1137,7 @@ lemma not_nb[simp]: "bound0 p \ bound0 (not p)" by (induct p, auto) -constdefs conj :: "fm \ fm \ fm" +definition conj :: "fm \ fm \ fm" where "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else if p = q then p else And p q)" lemma conj[simp]: "Ifm bs (conj p q) = Ifm bs (And p q)" @@ -1148,7 +1148,7 @@ lemma conj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (conj p q)" using conj_def by auto -constdefs disj :: "fm \ fm \ fm" +definition disj :: "fm \ fm \ fm" where "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p else if p=q then p else Or p q)" @@ -1159,7 +1159,7 @@ lemma disj_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (disj p q)" using disj_def by auto -constdefs imp :: "fm \ fm \ fm" +definition imp :: "fm \ fm \ fm" where "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p else Imp p q)" lemma imp[simp]: "Ifm bs (imp p q) = Ifm bs (Imp p q)" @@ -1169,7 +1169,7 @@ lemma imp_nb[simp]: "\bound0 p ; bound0 q\ \ bound0 (imp p q)" using imp_def by (cases "p=F \ q=T \ p=q",simp_all add: imp_def) -constdefs iff :: "fm \ fm \ fm" +definition iff :: "fm \ fm \ fm" where "iff p q \ (if (p = q) then T else if (p = not q \ not p = q) then F else if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else Iff p q)" @@ -1216,7 +1216,7 @@ thus "real d rdvd real c * t" using d rdvd_mult[OF gnz, where n="d div g" and x="real (c div g) * t"] real_of_int_div[OF gnz gd] real_of_int_div[OF gnz gc] by simp qed -constdefs simpdvd:: "int \ num \ (int \ num)" +definition simpdvd :: "int \ num \ (int \ num)" where "simpdvd d t \ (let g = numgcd t in if g > 1 then (let g' = gcd d g in @@ -1479,7 +1479,7 @@ (* Generic quantifier elimination *) -constdefs list_conj :: "fm list \ fm" +definition list_conj :: "fm list \ fm" where "list_conj ps \ foldr conj ps T" lemma list_conj: "Ifm bs (list_conj ps) = (\p\ set ps. Ifm bs p)" by (induct ps, auto simp add: list_conj_def) @@ -1487,7 +1487,7 @@ by (induct ps, auto simp add: list_conj_def) lemma list_conj_nb: " \p\ set ps. bound0 p \ bound0 (list_conj ps)" by (induct ps, auto simp add: list_conj_def) -constdefs CJNB:: "(fm \ fm) \ fm \ fm" +definition CJNB :: "(fm \ fm) \ fm \ fm" where "CJNB f p \ (let cjs = conjuncts p ; (yes,no) = List.partition bound0 cjs in conj (decr (list_conj yes)) (f (list_conj no)))" @@ -2954,7 +2954,7 @@ else (NDvd (i*k) (CN 0 c (Mul k e))))" "a\ p = (\ k. p)" -constdefs \ :: "fm \ int \ num \ fm" +definition \ :: "fm \ int \ num \ fm" where "\ p k t \ And (Dvd k t) (\\ p (t,k))" lemma \\: @@ -3517,7 +3517,7 @@ "isrlfm (Ge (CN 0 c e)) = (c>0 \ numbound0 e)" "isrlfm p = (isatom p \ (bound0 p))" -constdefs fp :: "fm \ int \ num \ int \ fm" +definition fp :: "fm \ int \ num \ int \ fm" where "fp p n s j \ (if n > 0 then (And p (And (Ge (CN 0 n (Sub s (Add (Floor s) (C j))))) (Lt (CN 0 n (Sub s (Add (Floor s) (C (j+1)))))))) @@ -3836,7 +3836,7 @@ (* Linearize a formula where Bound 0 ranges over [0,1) *) -constdefs rsplit :: "(int \ num \ fm) \ num \ fm" +definition rsplit :: "(int \ num \ fm) \ num \ fm" where "rsplit f a \ foldr disj (map (\ (\, n, s). conj \ (f n s)) (rsplit0 a)) F" lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\ x \ set xs. Ifm bs (f x))" @@ -5103,7 +5103,7 @@ (* Implement the right hand sides of Cooper's theorem and Ferrante and Rackoff. *) -constdefs ferrack01:: "fm \ fm" +definition ferrack01 :: "fm \ fm" where "ferrack01 p \ (let p' = rlfm(And (And (Ge(CN 0 1 (C 0))) (Lt (CN 0 1 (C (- 1))))) p); U = remdups(map simp_num_pair (map (\ ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m)) @@ -5350,7 +5350,7 @@ shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. d}. Ifm (real j#bs) (minusinf p)) \ (\ j\ {1.. d}. \ b\ (Inum (real i#bs)) ` set (\ p). Ifm ((b+real j)#bs) p))" using cp_thm[OF lp up dd dp] by auto -constdefs unit:: "fm \ fm \ num list \ int" +definition unit :: "fm \ fm \ num list \ int" where "unit p \ (let p' = zlfm p ; l = \ p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\ p' l); d = \ q; B = remdups (map simpnum (\ q)) ; a = remdups (map simpnum (\ q)) in if length B \ length a then (q,B,d) else (mirror q, a,d))" @@ -5417,7 +5417,7 @@ qed (* Cooper's Algorithm *) -constdefs cooper :: "fm \ fm" +definition cooper :: "fm \ fm" where "cooper p \ (let (q,B,d) = unit p; js = iupt (1,d); mq = simpfm (minusinf q); @@ -5561,7 +5561,7 @@ shows "(\ (x::int). Ifm (real x#bs) p) = ((\ j\ {1 .. \ p}. Ifm (real j#bs) (minusinf p)) \ (\ (e,c) \ R. \ j\ {1.. c*(\ p)}. Ifm (a#bs) (\ p c (Add e (C j)))))" using rl_thm[OF lp] \_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp -constdefs chooset:: "fm \ fm \ ((num\int) list) \ int" +definition chooset :: "fm \ fm \ ((num\int) list) \ int" where "chooset p \ (let q = zlfm p ; d = \ q; B = remdups (map (\ (t,k). (simpnum t,k)) (\ q)) ; a = remdups (map (\ (t,k). (simpnum t,k)) (\\ q)) @@ -5621,7 +5621,7 @@ ultimately show ?thes by blast qed -constdefs stage:: "fm \ int \ (num \ int) \ fm" +definition stage :: "fm \ int \ (num \ int) \ fm" where "stage p d \ (\ (e,c). evaldjf (\ j. simpfm (\ p c (Add e (C j)))) (iupt (1,c*d)))" lemma stage: shows "Ifm bs (stage p d (e,c)) = (\ j\{1 .. c*d}. Ifm bs (\ p c (Add e (C j))))" @@ -5641,7 +5641,7 @@ from evaldjf_bound0[OF th] show ?thesis by (unfold stage_def split_def) simp qed -constdefs redlove:: "fm \ fm" +definition redlove :: "fm \ fm" where "redlove p \ (let (q,B,d) = chooset p; mq = simpfm (minusinf q); diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy --- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Mar 01 13:42:31 2010 +0100 @@ -273,10 +273,10 @@ assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})" shows "allpolys isnpoly t \ isnpoly c \ allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm) -constdefs tmneg :: "tm \ tm" +definition tmneg :: "tm \ tm" where "tmneg t \ tmmul t (C (- 1,1))" -constdefs tmsub :: "tm \ tm \ tm" +definition tmsub :: "tm \ tm \ tm" where "tmsub s t \ (if s = t then CP 0\<^sub>p else tmadd (s,tmneg t))" lemma tmneg[simp]: "Itm vs bs (tmneg t) = Itm vs bs (Neg t)" @@ -477,26 +477,26 @@ lemma not[simp]: "Ifm vs bs (not p) = Ifm vs bs (NOT p)" by (induct p rule: not.induct) auto -constdefs conj :: "fm \ fm \ fm" +definition conj :: "fm \ fm \ fm" where "conj p q \ (if (p = F \ q=F) then F else if p=T then q else if q=T then p else if p = q then p else And p q)" lemma conj[simp]: "Ifm vs bs (conj p q) = Ifm vs bs (And p q)" by (cases "p=F \ q=F",simp_all add: conj_def) (cases p,simp_all) -constdefs disj :: "fm \ fm \ fm" +definition disj :: "fm \ fm \ fm" where "disj p q \ (if (p = T \ q=T) then T else if p=F then q else if q=F then p else if p=q then p else Or p q)" lemma disj[simp]: "Ifm vs bs (disj p q) = Ifm vs bs (Or p q)" by (cases "p=T \ q=T",simp_all add: disj_def) (cases p,simp_all) -constdefs imp :: "fm \ fm \ fm" +definition imp :: "fm \ fm \ fm" where "imp p q \ (if (p = F \ q=T \ p=q) then T else if p=T then q else if q=F then not p else Imp p q)" lemma imp[simp]: "Ifm vs bs (imp p q) = Ifm vs bs (Imp p q)" by (cases "p=F \ q=T",simp_all add: imp_def) -constdefs iff :: "fm \ fm \ fm" +definition iff :: "fm \ fm \ fm" where "iff p q \ (if (p = q) then T else if (p = NOT q \ NOT p = q) then F else if p=F then not q else if q=F then not p else if p=T then q else if q=T then p else Iff p q)" @@ -776,10 +776,10 @@ lemma bound0_qf: "bound0 p \ qfree p" by (induct p, simp_all) -constdefs djf:: "('a \ fm) \ 'a \ fm \ fm" +definition djf :: "('a \ fm) \ 'a \ fm \ fm" where "djf f p q \ (if q=T then T else if q=F then f p else (let fp = f p in case fp of T \ T | F \ q | _ \ Or (f p) q))" -constdefs evaldjf:: "('a \ fm) \ 'a list \ fm" +definition evaldjf :: "('a \ fm) \ 'a list \ fm" where "evaldjf f ps \ foldr (djf f) ps F" lemma djf_Or: "Ifm vs bs (djf f p q) = Ifm vs bs (Or (f p) q)" @@ -823,7 +823,7 @@ thus ?thesis by (simp only: list_all_iff) qed -constdefs DJ :: "(fm \ fm) \ fm \ fm" +definition DJ :: "(fm \ fm) \ fm \ fm" where "DJ f p \ evaldjf f (disjuncts p)" lemma DJ: assumes fdj: "\ p q. Ifm vs bs (f (Or p q)) = Ifm vs bs (Or (f p) (f q))" @@ -869,10 +869,10 @@ "conjuncts T = []" "conjuncts p = [p]" -constdefs list_conj :: "fm list \ fm" +definition list_conj :: "fm list \ fm" where "list_conj ps \ foldr conj ps T" -constdefs CJNB:: "(fm \ fm) \ fm \ fm" +definition CJNB :: "(fm \ fm) \ fm \ fm" where "CJNB f p \ (let cjs = conjuncts p ; (yes,no) = partition bound0 cjs in conj (decr0 (list_conj yes)) (f (list_conj no)))" @@ -1158,7 +1158,7 @@ "conjs p = [p]" lemma conjs_ci: "(\ q \ set (conjs p). Ifm vs bs q) = Ifm vs bs p" by (induct p rule: conjs.induct, auto) -constdefs list_disj :: "fm list \ fm" +definition list_disj :: "fm list \ fm" where "list_disj ps \ foldr disj ps F" lemma list_conj: "Ifm vs bs (list_conj ps) = (\p\ set ps. Ifm vs bs p)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy --- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Mon Mar 01 13:42:31 2010 +0100 @@ -188,12 +188,12 @@ | "poly_cmul y (CN c n p) = CN (poly_cmul y c) n (poly_cmul y p)" | "poly_cmul y p = C y *\<^sub>p p" -constdefs monic:: "poly \ (poly \ bool)" +definition monic :: "poly \ (poly \ bool)" where "monic p \ (let h = headconst p in if h = 0\<^sub>N then (p,False) else ((C (Ninv h)) *\<^sub>p p, 0>\<^sub>N h))" subsection{* Pseudo-division *} -constdefs shift1:: "poly \ poly" +definition shift1 :: "poly \ poly" where "shift1 p \ CN 0\<^sub>p 0 p" consts funpow :: "nat \ ('a \ 'a) \ 'a \ 'a" @@ -212,7 +212,7 @@ by pat_completeness auto -constdefs polydivide:: "poly \ poly \ (nat \ poly)" +definition polydivide :: "poly \ poly \ (nat \ poly)" where "polydivide s p \ polydivide_aux (head p,degree p,p,0, s)" fun poly_deriv_aux :: "nat \ poly \ poly" where @@ -262,7 +262,7 @@ lemma isnpolyh_mono: "\n' \ n ; isnpolyh p n\ \ isnpolyh p n'" by (induct p rule: isnpolyh.induct, auto) -constdefs isnpoly:: "poly \ bool" +definition isnpoly :: "poly \ bool" where "isnpoly p \ isnpolyh p 0" text{* polyadd preserves normal forms *} diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Finite_Set.thy Mon Mar 01 13:42:31 2010 +0100 @@ -2528,8 +2528,7 @@ fold1Set_insertI [intro]: "\ fold_graph f a A x; a \ A \ \ fold1Set f (insert a A) x" -constdefs - fold1 :: "('a => 'a => 'a) => 'a set => 'a" +definition fold1 :: "('a => 'a => 'a) => 'a set => 'a" where "fold1 f A == THE x. fold1Set f A x" lemma fold1Set_nonempty: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Fun.thy --- a/src/HOL/Fun.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Fun.thy Mon Mar 01 13:42:31 2010 +0100 @@ -119,8 +119,9 @@ subsection {* Injectivity and Surjectivity *} -constdefs - inj_on :: "['a => 'b, 'a set] => bool" -- "injective" +definition + inj_on :: "['a => 'b, 'a set] => bool" where + -- "injective" "inj_on f A == ! x:A. ! y:A. f(x)=f(y) --> x=y" text{*A common special case: functions injective over the entire domain type.*} @@ -132,11 +133,14 @@ bij_betw :: "('a => 'b) => 'a set => 'b set => bool" where -- "bijective" [code del]: "bij_betw f A B \ inj_on f A & f ` A = B" -constdefs - surj :: "('a => 'b) => bool" (*surjective*) +definition + surj :: "('a => 'b) => bool" where + -- "surjective" "surj f == ! y. ? x. y=f(x)" - bij :: "('a => 'b) => bool" (*bijective*) +definition + bij :: "('a => 'b) => bool" where + -- "bijective" "bij f == inj f & surj f" lemma injI: @@ -377,8 +381,8 @@ subsection{*Function Updating*} -constdefs - fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" +definition + fun_upd :: "('a => 'b) => 'a => 'b => ('a => 'b)" where "fun_upd f a b == % x. if x=a then b else f x" nonterminals diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/HOL.thy --- a/src/HOL/HOL.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/HOL.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1122,8 +1122,7 @@ its premise. *} -constdefs - simp_implies :: "[prop, prop] => prop" (infixr "=simp=>" 1) +definition simp_implies :: "[prop, prop] => prop" (infixr "=simp=>" 1) where [code del]: "simp_implies \ op ==>" lemma simp_impliesI: @@ -1396,13 +1395,23 @@ ) *} -constdefs - induct_forall where "induct_forall P == \x. P x" - induct_implies where "induct_implies A B == A \ B" - induct_equal where "induct_equal x y == x = y" - induct_conj where "induct_conj A B == A \ B" - induct_true where "induct_true == True" - induct_false where "induct_false == False" +definition induct_forall where + "induct_forall P == \x. P x" + +definition induct_implies where + "induct_implies A B == A \ B" + +definition induct_equal where + "induct_equal x y == x = y" + +definition induct_conj where + "induct_conj A B == A \ B" + +definition induct_true where + "induct_true == True" + +definition induct_false where + "induct_false == False" lemma induct_forall_eq: "(!!x. P x) == Trueprop (induct_forall (\x. P x))" by (unfold atomize_all induct_forall_def) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hilbert_Choice.thy --- a/src/HOL/Hilbert_Choice.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hilbert_Choice.thy Mon Mar 01 13:42:31 2010 +0100 @@ -307,8 +307,8 @@ subsection {* Least value operator *} -constdefs - LeastM :: "['a => 'b::ord, 'a => bool] => 'a" +definition + LeastM :: "['a => 'b::ord, 'a => bool] => 'a" where "LeastM m P == SOME x. P x & (\y. P y --> m x <= m y)" syntax @@ -360,11 +360,12 @@ subsection {* Greatest value operator *} -constdefs - GreatestM :: "['a => 'b::ord, 'a => bool] => 'a" +definition + GreatestM :: "['a => 'b::ord, 'a => bool] => 'a" where "GreatestM m P == SOME x. P x & (\y. P y --> m y <= m x)" - Greatest :: "('a::ord => bool) => 'a" (binder "GREATEST " 10) +definition + Greatest :: "('a::ord => bool) => 'a" (binder "GREATEST " 10) where "Greatest == GreatestM (%x. x)" syntax diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare/Arith2.thy --- a/src/HOL/Hoare/Arith2.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare/Arith2.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Hoare/Arith2.thy - ID: $Id$ Author: Norbert Galm Copyright 1995 TUM @@ -10,11 +9,10 @@ imports Main begin -constdefs - "cd" :: "[nat, nat, nat] => bool" +definition "cd" :: "[nat, nat, nat] => bool" where "cd x m n == x dvd m & x dvd n" - gcd :: "[nat, nat] => nat" +definition gcd :: "[nat, nat] => nat" where "gcd m n == @x.(cd x m n) & (!y.(cd y m n) --> y<=x)" consts fac :: "nat => nat" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare/Heap.thy --- a/src/HOL/Hoare/Heap.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare/Heap.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Hoare/Heap.thy - ID: $Id$ Author: Tobias Nipkow Copyright 2002 TUM @@ -19,19 +18,17 @@ lemma not_Ref_eq [iff]: "(ALL y. x ~= Ref y) = (x = Null)" by (induct x) auto -consts addr :: "'a ref \ 'a" -primrec "addr(Ref a) = a" +primrec addr :: "'a ref \ 'a" where + "addr (Ref a) = a" section "The heap" subsection "Paths in the heap" -consts - Path :: "('a \ 'a ref) \ 'a ref \ 'a list \ 'a ref \ bool" -primrec -"Path h x [] y = (x = y)" -"Path h x (a#as) y = (x = Ref a \ Path h (h a) as y)" +primrec Path :: "('a \ 'a ref) \ 'a ref \ 'a list \ 'a ref \ bool" where + "Path h x [] y \ x = y" +| "Path h x (a#as) y \ x = Ref a \ Path h (h a) as y" lemma [iff]: "Path h Null xs y = (xs = [] \ y = Null)" apply(case_tac xs) @@ -60,8 +57,7 @@ subsection "Non-repeating paths" -constdefs - distPath :: "('a \ 'a ref) \ 'a ref \ 'a list \ 'a ref \ bool" +definition distPath :: "('a \ 'a ref) \ 'a ref \ 'a list \ 'a ref \ bool" where "distPath h x as y \ Path h x as y \ distinct as" text{* The term @{term"distPath h x as y"} expresses the fact that a @@ -86,8 +82,7 @@ subsubsection "Relational abstraction" -constdefs - List :: "('a \ 'a ref) \ 'a ref \ 'a list \ bool" +definition List :: "('a \ 'a ref) \ 'a ref \ 'a list \ bool" where "List h x as == Path h x as Null" lemma [simp]: "List h x [] = (x = Null)" @@ -138,10 +133,10 @@ subsection "Functional abstraction" -constdefs - islist :: "('a \ 'a ref) \ 'a ref \ bool" +definition islist :: "('a \ 'a ref) \ 'a ref \ bool" where "islist h p == \as. List h p as" - list :: "('a \ 'a ref) \ 'a ref \ 'a list" + +definition list :: "('a \ 'a ref) \ 'a ref \ 'a list" where "list h p == SOME as. List h p as" lemma List_conv_islist_list: "List h p as = (islist h p \ as = list h p)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare/Hoare_Logic.thy --- a/src/HOL/Hoare/Hoare_Logic.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare/Hoare_Logic.thy Mon Mar 01 13:42:31 2010 +0100 @@ -40,7 +40,7 @@ (s ~: b --> Sem c2 s s'))" "Sem(While b x c) s s' = (? n. iter n b (Sem c) s s')" -constdefs Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" +definition Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" where "Valid p c q == !s s'. Sem c s s' --> s : p --> s' : q" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare/Hoare_Logic_Abort.thy --- a/src/HOL/Hoare/Hoare_Logic_Abort.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy Mon Mar 01 13:42:31 2010 +0100 @@ -42,7 +42,7 @@ "Sem(While b x c) s s' = (if s = None then s' = None else \n. iter n b (Sem c) s s')" -constdefs Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" +definition Valid :: "'a bexp \ 'a com \ 'a bexp \ bool" where "Valid p c q == \s s'. Sem c s s' \ s : Some ` p \ s' : Some ` q" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare/Pointer_Examples.thy --- a/src/HOL/Hoare/Pointer_Examples.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare/Pointer_Examples.thy Mon Mar 01 13:42:31 2010 +0100 @@ -216,10 +216,10 @@ text"This is still a bit rough, especially the proof." -constdefs - cor :: "bool \ bool \ bool" +definition cor :: "bool \ bool \ bool" where "cor P Q == if P then True else Q" - cand :: "bool \ bool \ bool" + +definition cand :: "bool \ bool \ bool" where "cand P Q == if P then Q else False" consts merge :: "'a list * 'a list * ('a \ 'a \ bool) \ 'a list" @@ -481,7 +481,7 @@ subsection "Storage allocation" -constdefs new :: "'a set \ 'a" +definition new :: "'a set \ 'a" where "new A == SOME a. a \ A" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare/Pointers0.thy --- a/src/HOL/Hoare/Pointers0.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare/Pointers0.thy Mon Mar 01 13:42:31 2010 +0100 @@ -73,8 +73,7 @@ subsubsection "Relational abstraction" -constdefs - List :: "('a::ref \ 'a) \ 'a \ 'a list \ bool" +definition List :: "('a::ref \ 'a) \ 'a \ 'a list \ bool" where "List h x as == Path h x as Null" lemma [simp]: "List h x [] = (x = Null)" @@ -122,10 +121,10 @@ subsection "Functional abstraction" -constdefs - islist :: "('a::ref \ 'a) \ 'a \ bool" +definition islist :: "('a::ref \ 'a) \ 'a \ bool" where "islist h p == \as. List h p as" - list :: "('a::ref \ 'a) \ 'a \ 'a list" + +definition list :: "('a::ref \ 'a) \ 'a \ 'a list" where "list h p == SOME as. List h p as" lemma List_conv_islist_list: "List h p as = (islist h p \ as = list h p)" @@ -407,7 +406,7 @@ subsection "Storage allocation" -constdefs new :: "'a set \ 'a::ref" +definition new :: "'a set \ 'a::ref" where "new A == SOME a. a \ A & a \ Null" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare/SepLogHeap.thy --- a/src/HOL/Hoare/SepLogHeap.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare/SepLogHeap.thy Mon Mar 01 13:42:31 2010 +0100 @@ -41,8 +41,7 @@ subsection "Lists on the heap" -constdefs - List :: "heap \ nat \ nat list \ bool" +definition List :: "heap \ nat \ nat list \ bool" where "List h x as == Path h x as 0" lemma [simp]: "List h x [] = (x = 0)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare/Separation.thy --- a/src/HOL/Hoare/Separation.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare/Separation.thy Mon Mar 01 13:42:31 2010 +0100 @@ -16,20 +16,19 @@ text{* The semantic definition of a few connectives: *} -constdefs - ortho:: "heap \ heap \ bool" (infix "\" 55) +definition ortho :: "heap \ heap \ bool" (infix "\" 55) where "h1 \ h2 == dom h1 \ dom h2 = {}" - is_empty :: "heap \ bool" +definition is_empty :: "heap \ bool" where "is_empty h == h = empty" - singl:: "heap \ nat \ nat \ bool" +definition singl:: "heap \ nat \ nat \ bool" where "singl h x y == dom h = {x} & h x = Some y" - star:: "(heap \ bool) \ (heap \ bool) \ (heap \ bool)" +definition star:: "(heap \ bool) \ (heap \ bool) \ (heap \ bool)" where "star P Q == \h. \h1 h2. h = h1++h2 \ h1 \ h2 \ P h1 \ Q h2" - wand:: "(heap \ bool) \ (heap \ bool) \ (heap \ bool)" +definition wand:: "(heap \ bool) \ (heap \ bool) \ (heap \ bool)" where "wand P Q == \h. \h'. h' \ h \ P h' \ Q(h++h')" text{*This is what assertions look like without any syntactic sugar: *} diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare_Parallel/Gar_Coll.thy --- a/src/HOL/Hoare_Parallel/Gar_Coll.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare_Parallel/Gar_Coll.thy Mon Mar 01 13:42:31 2010 +0100 @@ -31,8 +31,7 @@ under which the selected edge @{text "R"} and node @{text "T"} are valid: *} -constdefs - Mut_init :: "gar_coll_state \ bool" +definition Mut_init :: "gar_coll_state \ bool" where "Mut_init \ \ T \ Reach \E \ R < length \E \ T < length \M \" text {* \noindent For the mutator we @@ -40,14 +39,13 @@ @{text "\z"} is set to false if the mutator has already redirected an edge but has not yet colored the new target. *} -constdefs - Redirect_Edge :: "gar_coll_state ann_com" +definition Redirect_Edge :: "gar_coll_state ann_com" where "Redirect_Edge \ .{\Mut_init \ \z}. \\E:=\E[R:=(fst(\E!R), T)],, \z:= (\\z)\" - Color_Target :: "gar_coll_state ann_com" +definition Color_Target :: "gar_coll_state ann_com" where "Color_Target \ .{\Mut_init \ \\z}. \\M:=\M[T:=Black],, \z:= (\\z)\" - Mutator :: "gar_coll_state ann_com" +definition Mutator :: "gar_coll_state ann_com" where "Mutator \ .{\Mut_init \ \z}. WHILE True INV .{\Mut_init \ \z}. @@ -88,22 +86,20 @@ consts M_init :: nodes -constdefs - Proper_M_init :: "gar_coll_state \ bool" +definition Proper_M_init :: "gar_coll_state \ bool" where "Proper_M_init \ \ Blacks M_init=Roots \ length M_init=length \M \" - Proper :: "gar_coll_state \ bool" +definition Proper :: "gar_coll_state \ bool" where "Proper \ \ Proper_Roots \M \ Proper_Edges(\M, \E) \ \Proper_M_init \" - Safe :: "gar_coll_state \ bool" +definition Safe :: "gar_coll_state \ bool" where "Safe \ \ Reach \E \ Blacks \M \" lemmas collector_defs = Proper_M_init_def Proper_def Safe_def subsubsection {* Blackening the roots *} -constdefs - Blacken_Roots :: " gar_coll_state ann_com" +definition Blacken_Roots :: " gar_coll_state ann_com" where "Blacken_Roots \ .{\Proper}. \ind:=0;; @@ -133,13 +129,11 @@ subsubsection {* Propagating black *} -constdefs - PBInv :: "gar_coll_state \ nat \ bool" +definition PBInv :: "gar_coll_state \ nat \ bool" where "PBInv \ \ \ind. \obc < Blacks \M \ (\i BtoW (\E!i, \M) \ (\\z \ i=R \ (snd(\E!R)) = T \ (\r. ind \ r \ r < length \E \ BtoW(\E!r,\M))))\" -constdefs - Propagate_Black_aux :: "gar_coll_state ann_com" +definition Propagate_Black_aux :: "gar_coll_state ann_com" where "Propagate_Black_aux \ .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M}. \ind:=0;; @@ -214,14 +208,12 @@ subsubsection {* Refining propagating black *} -constdefs - Auxk :: "gar_coll_state \ bool" +definition Auxk :: "gar_coll_state \ bool" where "Auxk \ \\kM \ (\M!\k\Black \ \BtoW(\E!\ind, \M) \ \obcM \ (\\z \ \ind=R \ snd(\E!R)=T \ (\r. \ind rE \ BtoW(\E!r, \M))))\" -constdefs - Propagate_Black :: " gar_coll_state ann_com" +definition Propagate_Black :: " gar_coll_state ann_com" where "Propagate_Black \ .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M}. \ind:=0;; @@ -328,12 +320,10 @@ subsubsection {* Counting black nodes *} -constdefs - CountInv :: "gar_coll_state \ nat \ bool" +definition CountInv :: "gar_coll_state \ nat \ bool" where "CountInv \ \ \ind. {i. i \Ma!i=Black}\\bc \" -constdefs - Count :: " gar_coll_state ann_com" +definition Count :: " gar_coll_state ann_com" where "Count \ .{\Proper \ Roots\Blacks \M \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M @@ -398,12 +388,10 @@ Append_to_free2: "i \ Reach e \ n \ Reach (Append_to_free(i, e)) = ( n = i \ n \ Reach e)" -constdefs - AppendInv :: "gar_coll_state \ nat \ bool" +definition AppendInv :: "gar_coll_state \ nat \ bool" where "AppendInv \ \\ind. \iM. ind\i \ i\Reach \E \ \M!i=Black\" -constdefs - Append :: " gar_coll_state ann_com" +definition Append :: " gar_coll_state ann_com" where "Append \ .{\Proper \ Roots\Blacks \M \ \Safe}. \ind:=0;; @@ -444,8 +432,7 @@ subsubsection {* Correctness of the Collector *} -constdefs - Collector :: " gar_coll_state ann_com" +definition Collector :: " gar_coll_state ann_com" where "Collector \ .{\Proper}. WHILE True INV .{\Proper}. diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare_Parallel/Graph.thy --- a/src/HOL/Hoare_Parallel/Graph.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare_Parallel/Graph.thy Mon Mar 01 13:42:31 2010 +0100 @@ -13,20 +13,19 @@ consts Roots :: "nat set" -constdefs - Proper_Roots :: "nodes \ bool" +definition Proper_Roots :: "nodes \ bool" where "Proper_Roots M \ Roots\{} \ Roots \ {i. i edges) \ bool" +definition Proper_Edges :: "(nodes \ edges) \ bool" where "Proper_Edges \ (\(M,E). \i snd(E!i) nodes) \ bool" +definition BtoW :: "(edge \ nodes) \ bool" where "BtoW \ (\(e,M). (M!fst e)=Black \ (M!snd e)\Black)" - Blacks :: "nodes \ nat set" +definition Blacks :: "nodes \ nat set" where "Blacks M \ {i. i M!i=Black}" - Reach :: "edges \ nat set" +definition Reach :: "edges \ nat set" where "Reach E \ {x. (\path. 1 path!(length path - 1)\Roots \ x=path!0 \ (\ij x\Roots}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy --- a/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Mon Mar 01 13:42:31 2010 +0100 @@ -26,24 +26,23 @@ subsection {* The Mutators *} -constdefs - Mul_mut_init :: "mul_gar_coll_state \ nat \ bool" +definition Mul_mut_init :: "mul_gar_coll_state \ nat \ bool" where "Mul_mut_init \ \ \n. n=length \Muts \ (\iMuts!i)E \ T (\Muts!i)M) \" - Mul_Redirect_Edge :: "nat \ nat \ mul_gar_coll_state ann_com" +definition Mul_Redirect_Edge :: "nat \ nat \ mul_gar_coll_state ann_com" where "Mul_Redirect_Edge j n \ .{\Mul_mut_init n \ Z (\Muts!j)}. \IF T(\Muts!j) \ Reach \E THEN \E:= \E[R (\Muts!j):= (fst (\E!R(\Muts!j)), T (\Muts!j))] FI,, \Muts:= \Muts[j:= (\Muts!j) \Z:=False\]\" - Mul_Color_Target :: "nat \ nat \ mul_gar_coll_state ann_com" +definition Mul_Color_Target :: "nat \ nat \ mul_gar_coll_state ann_com" where "Mul_Color_Target j n \ .{\Mul_mut_init n \ \ Z (\Muts!j)}. \\M:=\M[T (\Muts!j):=Black],, \Muts:=\Muts[j:= (\Muts!j) \Z:=True\]\" - Mul_Mutator :: "nat \ nat \ mul_gar_coll_state ann_com" +definition Mul_Mutator :: "nat \ nat \ mul_gar_coll_state ann_com" where "Mul_Mutator j n \ .{\Mul_mut_init n \ Z (\Muts!j)}. WHILE True @@ -156,28 +155,25 @@ subsection {* The Collector *} -constdefs - Queue :: "mul_gar_coll_state \ nat" +definition Queue :: "mul_gar_coll_state \ nat" where "Queue \ \ length (filter (\i. \ Z i \ \M!(T i) \ Black) \Muts) \" consts M_init :: nodes -constdefs - Proper_M_init :: "mul_gar_coll_state \ bool" +definition Proper_M_init :: "mul_gar_coll_state \ bool" where "Proper_M_init \ \ Blacks M_init=Roots \ length M_init=length \M \" - Mul_Proper :: "mul_gar_coll_state \ nat \ bool" +definition Mul_Proper :: "mul_gar_coll_state \ nat \ bool" where "Mul_Proper \ \ \n. Proper_Roots \M \ Proper_Edges (\M, \E) \ \Proper_M_init \ n=length \Muts \" - Safe :: "mul_gar_coll_state \ bool" +definition Safe :: "mul_gar_coll_state \ bool" where "Safe \ \ Reach \E \ Blacks \M \" lemmas mul_collector_defs = Proper_M_init_def Mul_Proper_def Safe_def subsubsection {* Blackening Roots *} -constdefs - Mul_Blacken_Roots :: "nat \ mul_gar_coll_state ann_com" +definition Mul_Blacken_Roots :: "nat \ mul_gar_coll_state ann_com" where "Mul_Blacken_Roots n \ .{\Mul_Proper n}. \ind:=0;; @@ -208,16 +204,14 @@ subsubsection {* Propagating Black *} -constdefs - Mul_PBInv :: "mul_gar_coll_state \ bool" +definition Mul_PBInv :: "mul_gar_coll_state \ bool" where "Mul_PBInv \ \\Safe \ \obc\Blacks \M \ \l<\Queue \ (\i<\ind. \BtoW(\E!i,\M)) \ \l\\Queue\" - Mul_Auxk :: "mul_gar_coll_state \ bool" +definition Mul_Auxk :: "mul_gar_coll_state \ bool" where "Mul_Auxk \ \\l<\Queue \ \M!\k\Black \ \BtoW(\E!\ind, \M) \ \obc\Blacks \M\" -constdefs - Mul_Propagate_Black :: "nat \ mul_gar_coll_state ann_com" +definition Mul_Propagate_Black :: "nat \ mul_gar_coll_state ann_com" where "Mul_Propagate_Black n \ .{\Mul_Proper n \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M \ (\Safe \ \l\\Queue \ \obc\Blacks \M)}. @@ -296,11 +290,10 @@ subsubsection {* Counting Black Nodes *} -constdefs - Mul_CountInv :: "mul_gar_coll_state \ nat \ bool" - "Mul_CountInv \ \ \ind. {i. i \Ma!i=Black}\\bc \" +definition Mul_CountInv :: "mul_gar_coll_state \ nat \ bool" where + "Mul_CountInv \ \ \ind. {i. i \Ma!i=Black}\\bc \" - Mul_Count :: "nat \ mul_gar_coll_state ann_com" +definition Mul_Count :: "nat \ mul_gar_coll_state ann_com" where "Mul_Count n \ .{\Mul_Proper n \ Roots\Blacks \M \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M @@ -396,11 +389,10 @@ Append_to_free2: "i \ Reach e \ n \ Reach (Append_to_free(i, e)) = ( n = i \ n \ Reach e)" -constdefs - Mul_AppendInv :: "mul_gar_coll_state \ nat \ bool" +definition Mul_AppendInv :: "mul_gar_coll_state \ nat \ bool" where "Mul_AppendInv \ \ \ind. (\i. ind\i \ iM \ i\Reach \E \ \M!i=Black)\" - Mul_Append :: "nat \ mul_gar_coll_state ann_com" +definition Mul_Append :: "nat \ mul_gar_coll_state ann_com" where "Mul_Append n \ .{\Mul_Proper n \ Roots\Blacks \M \ \Safe}. \ind:=0;; @@ -438,8 +430,7 @@ subsubsection {* Collector *} -constdefs - Mul_Collector :: "nat \ mul_gar_coll_state ann_com" +definition Mul_Collector :: "nat \ mul_gar_coll_state ann_com" where "Mul_Collector n \ .{\Mul_Proper n}. WHILE True INV .{\Mul_Proper n}. diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare_Parallel/OG_Hoare.thy --- a/src/HOL/Hoare_Parallel/OG_Hoare.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy Mon Mar 01 13:42:31 2010 +0100 @@ -27,12 +27,12 @@ consts post :: "'a ann_triple_op \ 'a assn" primrec "post (c, q) = q" -constdefs interfree_aux :: "('a ann_com_op \ 'a assn \ 'a ann_com_op) \ bool" +definition interfree_aux :: "('a ann_com_op \ 'a assn \ 'a ann_com_op) \ bool" where "interfree_aux \ \(co, q, co'). co'= None \ (\(r,a) \ atomics (the co'). \= (q \ r) a q \ (co = None \ (\p \ assertions (the co). \= (p \ r) a p)))" -constdefs interfree :: "(('a ann_triple_op) list) \ bool" +definition interfree :: "(('a ann_triple_op) list) \ bool" where "interfree Ts \ \i j. i < length Ts \ j < length Ts \ i \ j \ interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) " diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare_Parallel/OG_Tactics.thy --- a/src/HOL/Hoare_Parallel/OG_Tactics.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy Mon Mar 01 13:42:31 2010 +0100 @@ -171,8 +171,7 @@ "\- (q \ (r \ b)) a q \ interfree_aux (None, q, Some (AnnAwait r b a))" by(simp add: interfree_aux_def oghoare_sound) -constdefs - interfree_swap :: "('a ann_triple_op * ('a ann_triple_op) list) \ bool" +definition interfree_swap :: "('a ann_triple_op * ('a ann_triple_op) list) \ bool" where "interfree_swap == \(x, xs). \y\set xs. interfree_aux (com x, post x, com y) \ interfree_aux(com y, post y, com x)" @@ -208,7 +207,7 @@ \ interfree (map (\k. (c k, Q k)) [a.. bool " ("[\] _" [0] 45) +definition map_ann_hoare :: "(('a ann_com_op * 'a assn) list) \ bool " ("[\] _" [0] 45) where "[\] Ts == (\ic q. Ts!i=(Some c, q) \ \ c q)" lemma MapAnnEmpty: "[\] []" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare_Parallel/OG_Tran.thy --- a/src/HOL/Hoare_Parallel/OG_Tran.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare_Parallel/OG_Tran.thy Mon Mar 01 13:42:31 2010 +0100 @@ -7,14 +7,13 @@ 'a ann_com_op = "('a ann_com) option" 'a ann_triple_op = "('a ann_com_op \ 'a assn)" -consts com :: "'a ann_triple_op \ 'a ann_com_op" -primrec "com (c, q) = c" +primrec com :: "'a ann_triple_op \ 'a ann_com_op" where + "com (c, q) = c" -consts post :: "'a ann_triple_op \ 'a assn" -primrec "post (c, q) = q" +primrec post :: "'a ann_triple_op \ 'a assn" where + "post (c, q) = q" -constdefs - All_None :: "'a ann_triple_op list \ bool" +definition All_None :: "'a ann_triple_op list \ bool" where "All_None Ts \ \(c, q) \ set Ts. c = None" subsection {* The Transition Relation *} @@ -88,26 +87,24 @@ subsection {* Definition of Semantics *} -constdefs - ann_sem :: "'a ann_com \ 'a \ 'a set" +definition ann_sem :: "'a ann_com \ 'a \ 'a set" where "ann_sem c \ \s. {t. (Some c, s) -*\ (None, t)}" - ann_SEM :: "'a ann_com \ 'a set \ 'a set" +definition ann_SEM :: "'a ann_com \ 'a set \ 'a set" where "ann_SEM c S \ \ann_sem c ` S" - sem :: "'a com \ 'a \ 'a set" +definition sem :: "'a com \ 'a \ 'a set" where "sem c \ \s. {t. \Ts. (c, s) -P*\ (Parallel Ts, t) \ All_None Ts}" - SEM :: "'a com \ 'a set \ 'a set" +definition SEM :: "'a com \ 'a set \ 'a set" where "SEM c S \ \sem c ` S " abbreviation Omega :: "'a com" ("\" 63) where "\ \ While UNIV UNIV (Basic id)" -consts fwhile :: "'a bexp \ 'a com \ nat \ 'a com" -primrec - "fwhile b c 0 = \" - "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)" +primrec fwhile :: "'a bexp \ 'a com \ nat \ 'a com" where + "fwhile b c 0 = \" + | "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)" subsubsection {* Proofs *} @@ -299,11 +296,10 @@ section {* Validity of Correctness Formulas *} -constdefs - com_validity :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\= _// _//_)" [90,55,90] 50) +definition com_validity :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\= _// _//_)" [90,55,90] 50) where "\= p c q \ SEM c p \ q" - ann_com_validity :: "'a ann_com \ 'a assn \ bool" ("\ _ _" [60,90] 45) +definition ann_com_validity :: "'a ann_com \ 'a assn \ bool" ("\ _ _" [60,90] 45) where "\ c q \ ann_SEM c (pre c) \ q" end \ No newline at end of file diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare_Parallel/RG_Hoare.thy --- a/src/HOL/Hoare_Parallel/RG_Hoare.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy Mon Mar 01 13:42:31 2010 +0100 @@ -7,8 +7,7 @@ declare Un_subset_iff [simp del] le_sup_iff [simp del] declare Cons_eq_map_conv [iff] -constdefs - stable :: "'a set \ ('a \ 'a) set \ bool" +definition stable :: "'a set \ ('a \ 'a) set \ bool" where "stable \ \f g. (\x y. x \ f \ (x, y) \ g \ y \ f)" inductive @@ -39,16 +38,19 @@ \ P sat [pre', rely', guar', post'] \ \ \ P sat [pre, rely, guar, post]" -constdefs - Pre :: "'a rgformula \ 'a set" +definition Pre :: "'a rgformula \ 'a set" where "Pre x \ fst(snd x)" - Post :: "'a rgformula \ 'a set" + +definition Post :: "'a rgformula \ 'a set" where "Post x \ snd(snd(snd(snd x)))" - Rely :: "'a rgformula \ ('a \ 'a) set" + +definition Rely :: "'a rgformula \ ('a \ 'a) set" where "Rely x \ fst(snd(snd x))" - Guar :: "'a rgformula \ ('a \ 'a) set" + +definition Guar :: "'a rgformula \ ('a \ 'a) set" where "Guar x \ fst(snd(snd(snd x)))" - Com :: "'a rgformula \ 'a com" + +definition Com :: "'a rgformula \ 'a com" where "Com x \ fst x" subsection {* Proof System for Parallel Programs *} @@ -1093,8 +1095,7 @@ subsection {* Soundness of the System for Parallel Programs *} -constdefs - ParallelCom :: "('a rgformula) list \ 'a par_com" +definition ParallelCom :: "('a rgformula) list \ 'a par_com" where "ParallelCom Ps \ map (Some \ fst) Ps" lemma two: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Hoare_Parallel/RG_Tran.thy --- a/src/HOL/Hoare_Parallel/RG_Tran.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Hoare_Parallel/RG_Tran.thy Mon Mar 01 13:42:31 2010 +0100 @@ -81,8 +81,7 @@ | CptnEnv: "(P, t)#xs \ cptn \ (P,s)#(P,t)#xs \ cptn" | CptnComp: "\(P,s) -c\ (Q,t); (Q, t)#xs \ cptn \ \ (P,s)#(Q,t)#xs \ cptn" -constdefs - cp :: "('a com) option \ 'a \ ('a confs) set" +definition cp :: "('a com) option \ 'a \ ('a confs) set" where "cp P s \ {l. l!0=(P,s) \ l \ cptn}" subsubsection {* Parallel computations *} @@ -95,14 +94,12 @@ | ParCptnEnv: "(P,t)#xs \ par_cptn \ (P,s)#(P,t)#xs \ par_cptn" | ParCptnComp: "\ (P,s) -pc\ (Q,t); (Q,t)#xs \ par_cptn \ \ (P,s)#(Q,t)#xs \ par_cptn" -constdefs - par_cp :: "'a par_com \ 'a \ ('a par_confs) set" +definition par_cp :: "'a par_com \ 'a \ ('a par_confs) set" where "par_cp P s \ {l. l!0=(P,s) \ l \ par_cptn}" subsection{* Modular Definition of Computation *} -constdefs - lift :: "'a com \ 'a conf \ 'a conf" +definition lift :: "'a com \ 'a conf \ 'a conf" where "lift Q \ \(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))" inductive_set cptn_mod :: "('a confs) set" @@ -380,38 +377,36 @@ types 'a rgformula = "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set" -constdefs - assum :: "('a set \ ('a \ 'a) set) \ ('a confs) set" +definition assum :: "('a set \ ('a \ 'a) set) \ ('a confs) set" where "assum \ \(pre, rely). {c. snd(c!0) \ pre \ (\i. Suc i c!i -e\ c!(Suc i) \ (snd(c!i), snd(c!Suc i)) \ rely)}" - comm :: "(('a \ 'a) set \ 'a set) \ ('a confs) set" +definition comm :: "(('a \ 'a) set \ 'a set) \ ('a confs) set" where "comm \ \(guar, post). {c. (\i. Suc i c!i -c\ c!(Suc i) \ (snd(c!i), snd(c!Suc i)) \ guar) \ (fst (last c) = None \ snd (last c) \ post)}" - com_validity :: "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" - ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) +definition com_validity :: "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" + ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) where "\ P sat [pre, rely, guar, post] \ \s. cp (Some P) s \ assum(pre, rely) \ comm(guar, post)" subsection {* Validity for Parallel Programs. *} -constdefs - All_None :: "(('a com) option) list \ bool" +definition All_None :: "(('a com) option) list \ bool" where "All_None xs \ \c\set xs. c=None" - par_assum :: "('a set \ ('a \ 'a) set) \ ('a par_confs) set" +definition par_assum :: "('a set \ ('a \ 'a) set) \ ('a par_confs) set" where "par_assum \ \(pre, rely). {c. snd(c!0) \ pre \ (\i. Suc i c!i -pe\ c!Suc i \ (snd(c!i), snd(c!Suc i)) \ rely)}" - par_comm :: "(('a \ 'a) set \ 'a set) \ ('a par_confs) set" +definition par_comm :: "(('a \ 'a) set \ 'a set) \ ('a par_confs) set" where "par_comm \ \(guar, post). {c. (\i. Suc i c!i -pc\ c!Suc i \ (snd(c!i), snd(c!Suc i)) \ guar) \ (All_None (fst (last c)) \ snd( last c) \ post)}" - par_com_validity :: "'a par_com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set -\ 'a set \ bool" ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) +definition par_com_validity :: "'a par_com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set +\ 'a set \ bool" ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) where "\ Ps SAT [pre, rely, guar, post] \ \s. par_cp Ps s \ par_assum(pre, rely) \ par_comm(guar, post)" @@ -419,23 +414,22 @@ subsubsection {* Definition of the conjoin operator *} -constdefs - same_length :: "'a par_confs \ ('a confs) list \ bool" +definition same_length :: "'a par_confs \ ('a confs) list \ bool" where "same_length c clist \ (\i ('a confs) list \ bool" +definition same_state :: "'a par_confs \ ('a confs) list \ bool" where "same_state c clist \ (\i j ('a confs) list \ bool" +definition same_program :: "'a par_confs \ ('a confs) list \ bool" where "same_program c clist \ (\jx. fst(nth x j)) clist)" - compat_label :: "'a par_confs \ ('a confs) list \ bool" +definition compat_label :: "'a par_confs \ ('a confs) list \ bool" where "compat_label c clist \ (\j. Suc j (c!j -pc\ c!Suc j \ (\i (clist!i)! Suc j \ (\li \ (clist!l)!j -e\ (clist!l)! Suc j))) \ (c!j -pe\ c!Suc j \ (\i (clist!i)! Suc j)))" - conjoin :: "'a par_confs \ ('a confs) list \ bool" ("_ \ _" [65,65] 64) +definition conjoin :: "'a par_confs \ ('a confs) list \ bool" ("_ \ _" [65,65] 64) where "c \ clist \ (same_length c clist) \ (same_state c clist) \ (same_program c clist) \ (compat_label c clist)" subsubsection {* Some previous lemmas *} diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/IOA/Solve.thy --- a/src/HOL/IOA/Solve.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/IOA/Solve.thy Mon Mar 01 13:42:31 2010 +0100 @@ -10,9 +10,7 @@ imports IOA begin -constdefs - - is_weak_pmap :: "['c => 'a, ('action,'c)ioa,('action,'a)ioa] => bool" +definition is_weak_pmap :: "['c => 'a, ('action,'c)ioa,('action,'a)ioa] => bool" where "is_weak_pmap f C A == (!s:starts_of(C). f(s):starts_of(A)) & (!s t a. reachable C s & diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Import/HOL/HOL4Base.thy --- a/src/HOL/Import/HOL/HOL4Base.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Import/HOL/HOL4Base.thy Mon Mar 01 13:42:31 2010 +0100 @@ -4,22 +4,19 @@ ;setup_theory bool -constdefs - ARB :: "'a" +definition ARB :: "'a" where "ARB == SOME x::'a::type. True" lemma ARB_DEF: "ARB = (SOME x::'a::type. True)" by (import bool ARB_DEF) -constdefs - IN :: "'a => ('a => bool) => bool" +definition IN :: "'a => ('a => bool) => bool" where "IN == %(x::'a::type) f::'a::type => bool. f x" lemma IN_DEF: "IN = (%(x::'a::type) f::'a::type => bool. f x)" by (import bool IN_DEF) -constdefs - RES_FORALL :: "('a => bool) => ('a => bool) => bool" +definition RES_FORALL :: "('a => bool) => ('a => bool) => bool" where "RES_FORALL == %(p::'a::type => bool) m::'a::type => bool. ALL x::'a::type. IN x p --> m x" @@ -28,8 +25,7 @@ ALL x::'a::type. IN x p --> m x)" by (import bool RES_FORALL_DEF) -constdefs - RES_EXISTS :: "('a => bool) => ('a => bool) => bool" +definition RES_EXISTS :: "('a => bool) => ('a => bool) => bool" where "RES_EXISTS == %(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x" @@ -37,8 +33,7 @@ (%(p::'a::type => bool) m::'a::type => bool. EX x::'a::type. IN x p & m x)" by (import bool RES_EXISTS_DEF) -constdefs - RES_EXISTS_UNIQUE :: "('a => bool) => ('a => bool) => bool" +definition RES_EXISTS_UNIQUE :: "('a => bool) => ('a => bool) => bool" where "RES_EXISTS_UNIQUE == %(p::'a::type => bool) m::'a::type => bool. RES_EXISTS p m & @@ -52,8 +47,7 @@ (%x::'a::type. RES_FORALL p (%y::'a::type. m x & m y --> x = y)))" by (import bool RES_EXISTS_UNIQUE_DEF) -constdefs - RES_SELECT :: "('a => bool) => ('a => bool) => 'a" +definition RES_SELECT :: "('a => bool) => ('a => bool) => 'a" where "RES_SELECT == %(p::'a::type => bool) m::'a::type => bool. SOME x::'a::type. IN x p & m x" @@ -264,15 +258,13 @@ ;setup_theory combin -constdefs - K :: "'a => 'b => 'a" +definition K :: "'a => 'b => 'a" where "K == %(x::'a::type) y::'b::type. x" lemma K_DEF: "K = (%(x::'a::type) y::'b::type. x)" by (import combin K_DEF) -constdefs - S :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c" +definition S :: "('a => 'b => 'c) => ('a => 'b) => 'a => 'c" where "S == %(f::'a::type => 'b::type => 'c::type) (g::'a::type => 'b::type) x::'a::type. f x (g x)" @@ -282,8 +274,7 @@ x::'a::type. f x (g x))" by (import combin S_DEF) -constdefs - I :: "'a => 'a" +definition I :: "'a => 'a" where "(op ==::('a::type => 'a::type) => ('a::type => 'a::type) => prop) (I::'a::type => 'a::type) ((S::('a::type => ('a::type => 'a::type) => 'a::type) @@ -299,16 +290,14 @@ (K::'a::type => 'a::type => 'a::type))" by (import combin I_DEF) -constdefs - C :: "('a => 'b => 'c) => 'b => 'a => 'c" +definition C :: "('a => 'b => 'c) => 'b => 'a => 'c" where "C == %(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x" lemma C_DEF: "C = (%(f::'a::type => 'b::type => 'c::type) (x::'b::type) y::'a::type. f y x)" by (import combin C_DEF) -constdefs - W :: "('a => 'a => 'b) => 'a => 'b" +definition W :: "('a => 'a => 'b) => 'a => 'b" where "W == %(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x" lemma W_DEF: "W = (%(f::'a::type => 'a::type => 'b::type) x::'a::type. f x x)" @@ -582,8 +571,7 @@ ;setup_theory relation -constdefs - TC :: "('a => 'a => bool) => 'a => 'a => bool" +definition TC :: "('a => 'a => bool) => 'a => 'a => bool" where "TC == %(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type. ALL P::'a::type => 'a::type => bool. @@ -601,8 +589,7 @@ P a b)" by (import relation TC_DEF) -constdefs - RTC :: "('a => 'a => bool) => 'a => 'a => bool" +definition RTC :: "('a => 'a => bool) => 'a => 'a => bool" where "RTC == %(R::'a::type => 'a::type => bool) (a::'a::type) b::'a::type. ALL P::'a::type => 'a::type => bool. @@ -644,8 +631,7 @@ (ALL (x::'a::type) (y::'a::type) z::'a::type. R x y & R y z --> R x z)" by (import relation transitive_def) -constdefs - pred_reflexive :: "('a => 'a => bool) => bool" +definition pred_reflexive :: "('a => 'a => bool) => bool" where "pred_reflexive == %R::'a::type => 'a::type => bool. ALL x::'a::type. R x x" lemma reflexive_def: "ALL R::'a::type => 'a::type => bool. @@ -788,8 +774,7 @@ (ALL (x::'a::type) y::'a::type. RTC R x y --> RTC Q x y)" by (import relation RTC_MONOTONE) -constdefs - WF :: "('a => 'a => bool) => bool" +definition WF :: "('a => 'a => bool) => bool" where "WF == %R::'a::type => 'a::type => bool. ALL B::'a::type => bool. @@ -814,8 +799,7 @@ WF x --> x xa xb --> xa ~= xb" by (import relation WF_NOT_REFL) -constdefs - EMPTY_REL :: "'a => 'a => bool" +definition EMPTY_REL :: "'a => 'a => bool" where "EMPTY_REL == %(x::'a::type) y::'a::type. False" lemma EMPTY_REL_DEF: "ALL (x::'a::type) y::'a::type. EMPTY_REL x y = False" @@ -847,8 +831,7 @@ WF R --> WF (relation.inv_image R f)" by (import relation WF_inv_image) -constdefs - RESTRICT :: "('a => 'b) => ('a => 'a => bool) => 'a => 'a => 'b" +definition RESTRICT :: "('a => 'b) => ('a => 'a => bool) => 'a => 'a => 'b" where "RESTRICT == %(f::'a::type => 'b::type) (R::'a::type => 'a::type => bool) (x::'a::type) y::'a::type. if R y x then f y else ARB" @@ -891,8 +874,7 @@ the_fun R M x = Eps (approx R M x)" by (import relation the_fun_def) -constdefs - WFREC :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'b" +definition WFREC :: "('a => 'a => bool) => (('a => 'b) => 'a => 'b) => 'a => 'b" where "WFREC == %(R::'a::type => 'a::type => bool) (M::('a::type => 'b::type) => 'a::type => 'b::type) x::'a::type. @@ -1052,8 +1034,7 @@ split xb x = split f' xa" by (import pair pair_case_cong) -constdefs - LEX :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" +definition LEX :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" where "LEX == %(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool) (s::'a::type, t::'b::type) (u::'a::type, v::'b::type). @@ -1069,8 +1050,7 @@ WF x & WF xa --> WF (LEX x xa)" by (import pair WF_LEX) -constdefs - RPROD :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" +definition RPROD :: "('a => 'a => bool) => ('b => 'b => bool) => 'a * 'b => 'a * 'b => bool" where "RPROD == %(R1::'a::type => 'a::type => bool) (R2::'b::type => 'b::type => bool) (s::'a::type, t::'b::type) (u::'a::type, v::'b::type). R1 s u & R2 t v" @@ -1113,8 +1093,7 @@ lemma NOT_LESS_EQ: "ALL (m::nat) n::nat. m = n --> ~ m < n" by (import prim_rec NOT_LESS_EQ) -constdefs - SIMP_REC_REL :: "(nat => 'a) => 'a => ('a => 'a) => nat => bool" +definition SIMP_REC_REL :: "(nat => 'a) => 'a => ('a => 'a) => nat => bool" where "(op ==::((nat => 'a::type) => 'a::type => ('a::type => 'a::type) => nat => bool) => ((nat => 'a::type) @@ -1187,8 +1166,7 @@ (ALL m::nat. SIMP_REC x f (Suc m) = f (SIMP_REC x f m))" by (import prim_rec SIMP_REC_THM) -constdefs - PRE :: "nat => nat" +definition PRE :: "nat => nat" where "PRE == %m::nat. if m = 0 then 0 else SOME n::nat. m = Suc n" lemma PRE_DEF: "ALL m::nat. PRE m = (if m = 0 then 0 else SOME n::nat. m = Suc n)" @@ -1197,8 +1175,7 @@ lemma PRE: "PRE 0 = 0 & (ALL m::nat. PRE (Suc m) = m)" by (import prim_rec PRE) -constdefs - PRIM_REC_FUN :: "'a => ('a => nat => 'a) => nat => nat => 'a" +definition PRIM_REC_FUN :: "'a => ('a => nat => 'a) => nat => nat => 'a" where "PRIM_REC_FUN == %(x::'a::type) f::'a::type => nat => 'a::type. SIMP_REC (%n::nat. x) (%(fun::nat => 'a::type) n::nat. f (fun (PRE n)) n)" @@ -1214,8 +1191,7 @@ PRIM_REC_FUN x f (Suc m) n = f (PRIM_REC_FUN x f m (PRE n)) n)" by (import prim_rec PRIM_REC_EQN) -constdefs - PRIM_REC :: "'a => ('a => nat => 'a) => nat => 'a" +definition PRIM_REC :: "'a => ('a => nat => 'a) => nat => 'a" where "PRIM_REC == %(x::'a::type) (f::'a::type => nat => 'a::type) m::nat. PRIM_REC_FUN x f m (PRE m)" @@ -1286,8 +1262,7 @@ ;setup_theory arithmetic -constdefs - nat_elim__magic :: "nat => nat" +definition nat_elim__magic :: "nat => nat" where "nat_elim__magic == %n::nat. n" lemma nat_elim__magic: "ALL n::nat. nat_elim__magic n = n" @@ -1746,22 +1721,19 @@ ;setup_theory hrat -constdefs - trat_1 :: "nat * nat" +definition trat_1 :: "nat * nat" where "trat_1 == (0, 0)" lemma trat_1: "trat_1 = (0, 0)" by (import hrat trat_1) -constdefs - trat_inv :: "nat * nat => nat * nat" +definition trat_inv :: "nat * nat => nat * nat" where "trat_inv == %(x::nat, y::nat). (y, x)" lemma trat_inv: "ALL (x::nat) y::nat. trat_inv (x, y) = (y, x)" by (import hrat trat_inv) -constdefs - trat_add :: "nat * nat => nat * nat => nat * nat" +definition trat_add :: "nat * nat => nat * nat => nat * nat" where "trat_add == %(x::nat, y::nat) (x'::nat, y'::nat). (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))" @@ -1771,8 +1743,7 @@ (PRE (Suc x * Suc y' + Suc x' * Suc y), PRE (Suc y * Suc y'))" by (import hrat trat_add) -constdefs - trat_mul :: "nat * nat => nat * nat => nat * nat" +definition trat_mul :: "nat * nat => nat * nat => nat * nat" where "trat_mul == %(x::nat, y::nat) (x'::nat, y'::nat). (PRE (Suc x * Suc x'), PRE (Suc y * Suc y'))" @@ -1788,8 +1759,7 @@ (ALL n::nat. trat_sucint (Suc n) = trat_add (trat_sucint n) trat_1)" by (import hrat trat_sucint) -constdefs - trat_eq :: "nat * nat => nat * nat => bool" +definition trat_eq :: "nat * nat => nat * nat => bool" where "trat_eq == %(x::nat, y::nat) (x'::nat, y'::nat). Suc x * Suc y' = Suc x' * Suc y" @@ -1901,23 +1871,20 @@ (EX x::nat * nat. r = trat_eq x) = (dest_hrat (mk_hrat r) = r))" by (import hrat hrat_tybij) -constdefs - hrat_1 :: "hrat" +definition hrat_1 :: "hrat" where "hrat_1 == mk_hrat (trat_eq trat_1)" lemma hrat_1: "hrat_1 = mk_hrat (trat_eq trat_1)" by (import hrat hrat_1) -constdefs - hrat_inv :: "hrat => hrat" +definition hrat_inv :: "hrat => hrat" where "hrat_inv == %T1::hrat. mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))" lemma hrat_inv: "ALL T1::hrat. hrat_inv T1 = mk_hrat (trat_eq (trat_inv (Eps (dest_hrat T1))))" by (import hrat hrat_inv) -constdefs - hrat_add :: "hrat => hrat => hrat" +definition hrat_add :: "hrat => hrat => hrat" where "hrat_add == %(T1::hrat) T2::hrat. mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))" @@ -1927,8 +1894,7 @@ mk_hrat (trat_eq (trat_add (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))" by (import hrat hrat_add) -constdefs - hrat_mul :: "hrat => hrat => hrat" +definition hrat_mul :: "hrat => hrat => hrat" where "hrat_mul == %(T1::hrat) T2::hrat. mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))" @@ -1938,8 +1904,7 @@ mk_hrat (trat_eq (trat_mul (Eps (dest_hrat T1)) (Eps (dest_hrat T2))))" by (import hrat hrat_mul) -constdefs - hrat_sucint :: "nat => hrat" +definition hrat_sucint :: "nat => hrat" where "hrat_sucint == %T1::nat. mk_hrat (trat_eq (trat_sucint T1))" lemma hrat_sucint: "ALL T1::nat. hrat_sucint T1 = mk_hrat (trat_eq (trat_sucint T1))" @@ -1987,8 +1952,7 @@ ;setup_theory hreal -constdefs - hrat_lt :: "hrat => hrat => bool" +definition hrat_lt :: "hrat => hrat => bool" where "hrat_lt == %(x::hrat) y::hrat. EX d::hrat. y = hrat_add x d" lemma hrat_lt: "ALL (x::hrat) y::hrat. hrat_lt x y = (EX d::hrat. y = hrat_add x d)" @@ -2096,8 +2060,7 @@ hrat_lt x y --> (EX xa::hrat. hrat_lt x xa & hrat_lt xa y)" by (import hreal HRAT_MEAN) -constdefs - isacut :: "(hrat => bool) => bool" +definition isacut :: "(hrat => bool) => bool" where "isacut == %C::hrat => bool. Ex C & @@ -2113,8 +2076,7 @@ (ALL x::hrat. C x --> (EX y::hrat. C y & hrat_lt x y)))" by (import hreal isacut) -constdefs - cut_of_hrat :: "hrat => hrat => bool" +definition cut_of_hrat :: "hrat => hrat => bool" where "cut_of_hrat == %(x::hrat) y::hrat. hrat_lt y x" lemma cut_of_hrat: "ALL x::hrat. cut_of_hrat x = (%y::hrat. hrat_lt y x)" @@ -2173,15 +2135,13 @@ (EX x::hrat. hreal.cut X x & ~ hreal.cut X (hrat_mul u x))" by (import hreal CUT_NEARTOP_MUL) -constdefs - hreal_1 :: "hreal" +definition hreal_1 :: "hreal" where "hreal_1 == hreal (cut_of_hrat hrat_1)" lemma hreal_1: "hreal_1 = hreal (cut_of_hrat hrat_1)" by (import hreal hreal_1) -constdefs - hreal_add :: "hreal => hreal => hreal" +definition hreal_add :: "hreal => hreal => hreal" where "hreal_add == %(X::hreal) Y::hreal. hreal @@ -2197,8 +2157,7 @@ w = hrat_add x y & hreal.cut X x & hreal.cut Y y)" by (import hreal hreal_add) -constdefs - hreal_mul :: "hreal => hreal => hreal" +definition hreal_mul :: "hreal => hreal => hreal" where "hreal_mul == %(X::hreal) Y::hreal. hreal @@ -2214,8 +2173,7 @@ w = hrat_mul x y & hreal.cut X x & hreal.cut Y y)" by (import hreal hreal_mul) -constdefs - hreal_inv :: "hreal => hreal" +definition hreal_inv :: "hreal => hreal" where "hreal_inv == %X::hreal. hreal @@ -2233,8 +2191,7 @@ (ALL x::hrat. hreal.cut X x --> hrat_lt (hrat_mul w x) d))" by (import hreal hreal_inv) -constdefs - hreal_sup :: "(hreal => bool) => hreal" +definition hreal_sup :: "(hreal => bool) => hreal" where "hreal_sup == %P::hreal => bool. hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)" @@ -2242,8 +2199,7 @@ hreal_sup P = hreal (%w::hrat. EX X::hreal. P X & hreal.cut X w)" by (import hreal hreal_sup) -constdefs - hreal_lt :: "hreal => hreal => bool" +definition hreal_lt :: "hreal => hreal => bool" where "hreal_lt == %(X::hreal) Y::hreal. X ~= Y & (ALL x::hrat. hreal.cut X x --> hreal.cut Y x)" @@ -2301,8 +2257,7 @@ lemma HREAL_NOZERO: "ALL (X::hreal) Y::hreal. hreal_add X Y ~= X" by (import hreal HREAL_NOZERO) -constdefs - hreal_sub :: "hreal => hreal => hreal" +definition hreal_sub :: "hreal => hreal => hreal" where "hreal_sub == %(Y::hreal) X::hreal. hreal @@ -2358,15 +2313,13 @@ (ALL x::nat. Suc (NUMERAL_BIT2 x) = NUMERAL_BIT1 (Suc x))" by (import numeral numeral_suc) -constdefs - iZ :: "nat => nat" +definition iZ :: "nat => nat" where "iZ == %x::nat. x" lemma iZ: "ALL x::nat. iZ x = x" by (import numeral iZ) -constdefs - iiSUC :: "nat => nat" +definition iiSUC :: "nat => nat" where "iiSUC == %n::nat. Suc (Suc n)" lemma iiSUC: "ALL n::nat. iiSUC n = Suc (Suc n)" @@ -2699,8 +2652,7 @@ iBIT_cases (NUMERAL_BIT2 n) zf bf1 bf2 = bf2 n)" by (import numeral iBIT_cases) -constdefs - iDUB :: "nat => nat" +definition iDUB :: "nat => nat" where "iDUB == %x::nat. x + x" lemma iDUB: "ALL x::nat. iDUB x = x + x" @@ -2771,8 +2723,7 @@ NUMERAL_BIT2 x * xa = iDUB (iZ (x * xa + xa))" by (import numeral numeral_mult) -constdefs - iSQR :: "nat => nat" +definition iSQR :: "nat => nat" where "iSQR == %x::nat. x * x" lemma iSQR: "ALL x::nat. iSQR x = x * x" @@ -2809,8 +2760,7 @@ ALL (xa::'A::type) y::'B::type. x (P xa y) = xa & Y (P xa y) = y)" by (import ind_type INJ_INVERSE2) -constdefs - NUMPAIR :: "nat => nat => nat" +definition NUMPAIR :: "nat => nat => nat" where "NUMPAIR == %(x::nat) y::nat. 2 ^ x * (2 * y + 1)" lemma NUMPAIR: "ALL (x::nat) y::nat. NUMPAIR x y = 2 ^ x * (2 * y + 1)" @@ -2831,8 +2781,7 @@ specification (NUMFST NUMSND) NUMPAIR_DEST: "ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & NUMSND (NUMPAIR x y) = y" by (import ind_type NUMPAIR_DEST) -constdefs - NUMSUM :: "bool => nat => nat" +definition NUMSUM :: "bool => nat => nat" where "NUMSUM == %(b::bool) x::nat. if b then Suc (2 * x) else 2 * x" lemma NUMSUM: "ALL (b::bool) x::nat. NUMSUM b x = (if b then Suc (2 * x) else 2 * x)" @@ -2849,8 +2798,7 @@ specification (NUMLEFT NUMRIGHT) NUMSUM_DEST: "ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & NUMRIGHT (NUMSUM x y) = y" by (import ind_type NUMSUM_DEST) -constdefs - INJN :: "nat => nat => 'a => bool" +definition INJN :: "nat => nat => 'a => bool" where "INJN == %(m::nat) (n::nat) a::'a::type. n = m" lemma INJN: "ALL m::nat. INJN m = (%(n::nat) a::'a::type. n = m)" @@ -2859,8 +2807,7 @@ lemma INJN_INJ: "ALL (n1::nat) n2::nat. (INJN n1 = INJN n2) = (n1 = n2)" by (import ind_type INJN_INJ) -constdefs - INJA :: "'a => nat => 'a => bool" +definition INJA :: "'a => nat => 'a => bool" where "INJA == %(a::'a::type) (n::nat) b::'a::type. b = a" lemma INJA: "ALL a::'a::type. INJA a = (%(n::nat) b::'a::type. b = a)" @@ -2869,8 +2816,7 @@ lemma INJA_INJ: "ALL (a1::'a::type) a2::'a::type. (INJA a1 = INJA a2) = (a1 = a2)" by (import ind_type INJA_INJ) -constdefs - INJF :: "(nat => nat => 'a => bool) => nat => 'a => bool" +definition INJF :: "(nat => nat => 'a => bool) => nat => 'a => bool" where "INJF == %(f::nat => nat => 'a::type => bool) n::nat. f (NUMFST n) (NUMSND n)" lemma INJF: "ALL f::nat => nat => 'a::type => bool. @@ -2881,8 +2827,7 @@ (INJF f1 = INJF f2) = (f1 = f2)" by (import ind_type INJF_INJ) -constdefs - INJP :: "(nat => 'a => bool) => (nat => 'a => bool) => nat => 'a => bool" +definition INJP :: "(nat => 'a => bool) => (nat => 'a => bool) => nat => 'a => bool" where "INJP == %(f1::nat => 'a::type => bool) (f2::nat => 'a::type => bool) (n::nat) a::'a::type. if NUMLEFT n then f1 (NUMRIGHT n) a else f2 (NUMRIGHT n) a" @@ -2898,8 +2843,7 @@ (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')" by (import ind_type INJP_INJ) -constdefs - ZCONSTR :: "nat => 'a => (nat => nat => 'a => bool) => nat => 'a => bool" +definition ZCONSTR :: "nat => 'a => (nat => nat => 'a => bool) => nat => 'a => bool" where "ZCONSTR == %(c::nat) (i::'a::type) r::nat => nat => 'a::type => bool. INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))" @@ -2908,8 +2852,7 @@ ZCONSTR c i r = INJP (INJN (Suc c)) (INJP (INJA i) (INJF r))" by (import ind_type ZCONSTR) -constdefs - ZBOT :: "nat => 'a => bool" +definition ZBOT :: "nat => 'a => bool" where "ZBOT == INJP (INJN 0) (SOME z::nat => 'a::type => bool. True)" lemma ZBOT: "ZBOT = INJP (INJN 0) (SOME z::nat => 'a::type => bool. True)" @@ -2919,8 +2862,7 @@ ZCONSTR x xa xb ~= ZBOT" by (import ind_type ZCONSTR_ZBOT) -constdefs - ZRECSPACE :: "(nat => 'a => bool) => bool" +definition ZRECSPACE :: "(nat => 'a => bool) => bool" where "ZRECSPACE == %a0::nat => 'a::type => bool. ALL ZRECSPACE'::(nat => 'a::type => bool) => bool. @@ -2993,15 +2935,13 @@ (ALL r::nat => 'a::type => bool. ZRECSPACE r = (dest_rec (mk_rec r) = r))" by (import ind_type recspace_repfns) -constdefs - BOTTOM :: "'a recspace" +definition BOTTOM :: "'a recspace" where "BOTTOM == mk_rec ZBOT" lemma BOTTOM: "BOTTOM = mk_rec ZBOT" by (import ind_type BOTTOM) -constdefs - CONSTR :: "nat => 'a => (nat => 'a recspace) => 'a recspace" +definition CONSTR :: "nat => 'a => (nat => 'a recspace) => 'a recspace" where "CONSTR == %(c::nat) (i::'a::type) r::nat => 'a::type recspace. mk_rec (ZCONSTR c i (%n::nat. dest_rec (r n)))" @@ -3049,15 +2989,13 @@ (ALL (a::'a::type) (f::nat => 'a::type) n::nat. FCONS a f (Suc n) = f n)" by (import ind_type FCONS) -constdefs - FNIL :: "nat => 'a" +definition FNIL :: "nat => 'a" where "FNIL == %n::nat. SOME x::'a::type. True" lemma FNIL: "ALL n::nat. FNIL n = (SOME x::'a::type. True)" by (import ind_type FNIL) -constdefs - ISO :: "('a => 'b) => ('b => 'a) => bool" +definition ISO :: "('a => 'b) => ('b => 'a) => bool" where "ISO == %(f::'a::type => 'b::type) g::'b::type => 'a::type. (ALL x::'b::type. f (g x) = x) & (ALL y::'a::type. g (f y) = y)" @@ -3434,8 +3372,7 @@ (EX x::'a::type. IN x s & (ALL y::'a::type. IN y s --> M x <= M y))" by (import pred_set SET_MINIMUM) -constdefs - EMPTY :: "'a => bool" +definition EMPTY :: "'a => bool" where "EMPTY == %x::'a::type. False" lemma EMPTY_DEF: "EMPTY = (%x::'a::type. False)" @@ -3468,8 +3405,7 @@ lemma EQ_UNIV: "(ALL x::'a::type. IN x (s::'a::type => bool)) = (s = pred_set.UNIV)" by (import pred_set EQ_UNIV) -constdefs - SUBSET :: "('a => bool) => ('a => bool) => bool" +definition SUBSET :: "('a => bool) => ('a => bool) => bool" where "SUBSET == %(s::'a::type => bool) t::'a::type => bool. ALL x::'a::type. IN x s --> IN x t" @@ -3501,8 +3437,7 @@ lemma UNIV_SUBSET: "ALL x::'a::type => bool. SUBSET pred_set.UNIV x = (x = pred_set.UNIV)" by (import pred_set UNIV_SUBSET) -constdefs - PSUBSET :: "('a => bool) => ('a => bool) => bool" +definition PSUBSET :: "('a => bool) => ('a => bool) => bool" where "PSUBSET == %(s::'a::type => bool) t::'a::type => bool. SUBSET s t & s ~= t" lemma PSUBSET_DEF: "ALL (s::'a::type => bool) t::'a::type => bool. @@ -3640,8 +3575,7 @@ pred_set.INTER (pred_set.UNION x xa) (pred_set.UNION x xb)" by (import pred_set INTER_OVER_UNION) -constdefs - DISJOINT :: "('a => bool) => ('a => bool) => bool" +definition DISJOINT :: "('a => bool) => ('a => bool) => bool" where "DISJOINT == %(s::'a::type => bool) t::'a::type => bool. pred_set.INTER s t = EMPTY" @@ -3672,8 +3606,7 @@ DISJOINT u (pred_set.UNION s t) = (DISJOINT s u & DISJOINT t u)" by (import pred_set DISJOINT_UNION_BOTH) -constdefs - DIFF :: "('a => bool) => ('a => bool) => 'a => bool" +definition DIFF :: "('a => bool) => ('a => bool) => 'a => bool" where "DIFF == %(s::'a::type => bool) t::'a::type => bool. GSPEC (%x::'a::type. (x, IN x s & ~ IN x t))" @@ -3702,8 +3635,7 @@ lemma DIFF_EQ_EMPTY: "ALL x::'a::type => bool. DIFF x x = EMPTY" by (import pred_set DIFF_EQ_EMPTY) -constdefs - INSERT :: "'a => ('a => bool) => 'a => bool" +definition INSERT :: "'a => ('a => bool) => 'a => bool" where "INSERT == %(x::'a::type) s::'a::type => bool. GSPEC (%y::'a::type. (y, y = x | IN y s))" @@ -3778,8 +3710,7 @@ DIFF (INSERT x s) t = (if IN x t then DIFF s t else INSERT x (DIFF s t))" by (import pred_set INSERT_DIFF) -constdefs - DELETE :: "('a => bool) => 'a => 'a => bool" +definition DELETE :: "('a => bool) => 'a => 'a => bool" where "DELETE == %(s::'a::type => bool) x::'a::type. DIFF s (INSERT x EMPTY)" lemma DELETE_DEF: "ALL (s::'a::type => bool) x::'a::type. DELETE s x = DIFF s (INSERT x EMPTY)" @@ -3852,8 +3783,7 @@ specification (CHOICE) CHOICE_DEF: "ALL x::'a::type => bool. x ~= EMPTY --> IN (CHOICE x) x" by (import pred_set CHOICE_DEF) -constdefs - REST :: "('a => bool) => 'a => bool" +definition REST :: "('a => bool) => 'a => bool" where "REST == %s::'a::type => bool. DELETE s (CHOICE s)" lemma REST_DEF: "ALL s::'a::type => bool. REST s = DELETE s (CHOICE s)" @@ -3871,8 +3801,7 @@ lemma REST_PSUBSET: "ALL x::'a::type => bool. x ~= EMPTY --> PSUBSET (REST x) x" by (import pred_set REST_PSUBSET) -constdefs - SING :: "('a => bool) => bool" +definition SING :: "('a => bool) => bool" where "SING == %s::'a::type => bool. EX x::'a::type. s = INSERT x EMPTY" lemma SING_DEF: "ALL s::'a::type => bool. SING s = (EX x::'a::type. s = INSERT x EMPTY)" @@ -3917,8 +3846,7 @@ lemma SING_IFF_EMPTY_REST: "ALL x::'a::type => bool. SING x = (x ~= EMPTY & REST x = EMPTY)" by (import pred_set SING_IFF_EMPTY_REST) -constdefs - IMAGE :: "('a => 'b) => ('a => bool) => 'b => bool" +definition IMAGE :: "('a => 'b) => ('a => bool) => 'b => bool" where "IMAGE == %(f::'a::type => 'b::type) s::'a::type => bool. GSPEC (%x::'a::type. (f x, IN x s))" @@ -3971,8 +3899,7 @@ (pred_set.INTER (IMAGE f s) (IMAGE f t))" by (import pred_set IMAGE_INTER) -constdefs - INJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" +definition INJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" where "INJ == %(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. (ALL x::'a::type. IN x s --> IN (f x) t) & @@ -3998,8 +3925,7 @@ (ALL xa::'a::type => bool. INJ x xa EMPTY = (xa = EMPTY))" by (import pred_set INJ_EMPTY) -constdefs - SURJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" +definition SURJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" where "SURJ == %(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. (ALL x::'a::type. IN x s --> IN (f x) t) & @@ -4028,8 +3954,7 @@ SURJ x xa xb = (IMAGE x xa = xb)" by (import pred_set IMAGE_SURJ) -constdefs - BIJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" +definition BIJ :: "('a => 'b) => ('a => bool) => ('b => bool) => bool" where "BIJ == %(f::'a::type => 'b::type) (s::'a::type => bool) t::'b::type => bool. INJ f s t & SURJ f s t" @@ -4065,8 +3990,7 @@ SURJ f s t --> (ALL x::'b::type. IN x t --> f (RINV f s x) = x)" by (import pred_set RINV_DEF) -constdefs - FINITE :: "('a => bool) => bool" +definition FINITE :: "('a => bool) => bool" where "FINITE == %s::'a::type => bool. ALL P::('a::type => bool) => bool. @@ -4219,8 +4143,7 @@ (ALL x::'a::type => bool. FINITE x --> P x)" by (import pred_set FINITE_COMPLETE_INDUCTION) -constdefs - INFINITE :: "('a => bool) => bool" +definition INFINITE :: "('a => bool) => bool" where "INFINITE == %s::'a::type => bool. ~ FINITE s" lemma INFINITE_DEF: "ALL s::'a::type => bool. INFINITE s = (~ FINITE s)" @@ -4320,8 +4243,7 @@ (f n)))))))))" by (import pred_set FINITE_WEAK_ENUMERATE) -constdefs - BIGUNION :: "(('a => bool) => bool) => 'a => bool" +definition BIGUNION :: "(('a => bool) => bool) => 'a => bool" where "BIGUNION == %P::('a::type => bool) => bool. GSPEC (%x::'a::type. (x, EX p::'a::type => bool. IN p P & IN x p))" @@ -4367,8 +4289,7 @@ FINITE (BIGUNION x)" by (import pred_set FINITE_BIGUNION) -constdefs - BIGINTER :: "(('a => bool) => bool) => 'a => bool" +definition BIGINTER :: "(('a => bool) => bool) => 'a => bool" where "BIGINTER == %B::('a::type => bool) => bool. GSPEC (%x::'a::type. (x, ALL P::'a::type => bool. IN P B --> IN x P))" @@ -4406,8 +4327,7 @@ DISJOINT x (BIGINTER xb) & DISJOINT (BIGINTER xb) x" by (import pred_set DISJOINT_BIGINTER) -constdefs - CROSS :: "('a => bool) => ('b => bool) => 'a * 'b => bool" +definition CROSS :: "('a => bool) => ('b => bool) => 'a * 'b => bool" where "CROSS == %(P::'a::type => bool) Q::'b::type => bool. GSPEC (%p::'a::type * 'b::type. (p, IN (fst p) P & IN (snd p) Q))" @@ -4460,8 +4380,7 @@ FINITE (CROSS P Q) = (P = EMPTY | Q = EMPTY | FINITE P & FINITE Q)" by (import pred_set FINITE_CROSS_EQ) -constdefs - COMPL :: "('a => bool) => 'a => bool" +definition COMPL :: "('a => bool) => 'a => bool" where "COMPL == DIFF pred_set.UNIV" lemma COMPL_DEF: "ALL P::'a::type => bool. COMPL P = DIFF pred_set.UNIV P" @@ -4513,8 +4432,7 @@ lemma CARD_COUNT: "ALL n::nat. CARD (count n) = n" by (import pred_set CARD_COUNT) -constdefs - ITSET_tupled :: "('a => 'b => 'b) => ('a => bool) * 'b => 'b" +definition ITSET_tupled :: "('a => 'b => 'b) => ('a => bool) * 'b => 'b" where "ITSET_tupled == %f::'a::type => 'b::type => 'b::type. WFREC @@ -4546,8 +4464,7 @@ else ARB)" by (import pred_set ITSET_tupled_primitive_def) -constdefs - ITSET :: "('a => 'b => 'b) => ('a => bool) => 'b => 'b" +definition ITSET :: "('a => 'b => 'b) => ('a => bool) => 'b => 'b" where "ITSET == %(f::'a::type => 'b::type => 'b::type) (x::'a::type => bool) x1::'b::type. ITSET_tupled f (x, x1)" @@ -4578,8 +4495,7 @@ ;setup_theory operator -constdefs - ASSOC :: "('a => 'a => 'a) => bool" +definition ASSOC :: "('a => 'a => 'a) => bool" where "ASSOC == %f::'a::type => 'a::type => 'a::type. ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z" @@ -4589,8 +4505,7 @@ (ALL (x::'a::type) (y::'a::type) z::'a::type. f x (f y z) = f (f x y) z)" by (import operator ASSOC_DEF) -constdefs - COMM :: "('a => 'a => 'b) => bool" +definition COMM :: "('a => 'a => 'b) => bool" where "COMM == %f::'a::type => 'a::type => 'b::type. ALL (x::'a::type) y::'a::type. f x y = f y x" @@ -4599,8 +4514,7 @@ COMM f = (ALL (x::'a::type) y::'a::type. f x y = f y x)" by (import operator COMM_DEF) -constdefs - FCOMM :: "('a => 'b => 'a) => ('c => 'a => 'a) => bool" +definition FCOMM :: "('a => 'b => 'a) => ('c => 'a => 'a) => bool" where "FCOMM == %(f::'a::type => 'b::type => 'a::type) g::'c::type => 'a::type => 'a::type. ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z" @@ -4611,8 +4525,7 @@ (ALL (x::'c::type) (y::'a::type) z::'b::type. g x (f y z) = f (g x y) z)" by (import operator FCOMM_DEF) -constdefs - RIGHT_ID :: "('a => 'b => 'a) => 'b => bool" +definition RIGHT_ID :: "('a => 'b => 'a) => 'b => bool" where "RIGHT_ID == %(f::'a::type => 'b::type => 'a::type) e::'b::type. ALL x::'a::type. f x e = x" @@ -4621,8 +4534,7 @@ RIGHT_ID f e = (ALL x::'a::type. f x e = x)" by (import operator RIGHT_ID_DEF) -constdefs - LEFT_ID :: "('a => 'b => 'b) => 'a => bool" +definition LEFT_ID :: "('a => 'b => 'b) => 'a => bool" where "LEFT_ID == %(f::'a::type => 'b::type => 'b::type) e::'a::type. ALL x::'b::type. f e x = x" @@ -4631,8 +4543,7 @@ LEFT_ID f e = (ALL x::'b::type. f e x = x)" by (import operator LEFT_ID_DEF) -constdefs - MONOID :: "('a => 'a => 'a) => 'a => bool" +definition MONOID :: "('a => 'a => 'a) => 'a => bool" where "MONOID == %(f::'a::type => 'a::type => 'a::type) e::'a::type. ASSOC f & RIGHT_ID f e & LEFT_ID f e" @@ -4690,15 +4601,13 @@ lemma IS_EL_DEF: "ALL (x::'a::type) l::'a::type list. x mem l = list_exists (op = x) l" by (import rich_list IS_EL_DEF) -constdefs - AND_EL :: "bool list => bool" +definition AND_EL :: "bool list => bool" where "AND_EL == list_all I" lemma AND_EL_DEF: "AND_EL = list_all I" by (import rich_list AND_EL_DEF) -constdefs - OR_EL :: "bool list => bool" +definition OR_EL :: "bool list => bool" where "OR_EL == list_exists I" lemma OR_EL_DEF: "OR_EL = list_exists I" @@ -4816,16 +4725,14 @@ (if P x then ([], x # l) else (x # fst (SPLITP P l), snd (SPLITP P l))))" by (import rich_list SPLITP) -constdefs - PREFIX :: "('a => bool) => 'a list => 'a list" +definition PREFIX :: "('a => bool) => 'a list => 'a list" where "PREFIX == %(P::'a::type => bool) l::'a::type list. fst (SPLITP (Not o P) l)" lemma PREFIX_DEF: "ALL (P::'a::type => bool) l::'a::type list. PREFIX P l = fst (SPLITP (Not o P) l)" by (import rich_list PREFIX_DEF) -constdefs - SUFFIX :: "('a => bool) => 'a list => 'a list" +definition SUFFIX :: "('a => bool) => 'a list => 'a list" where "SUFFIX == %P::'a::type => bool. foldl (%(l'::'a::type list) x::'a::type. if P x then SNOC x l' else []) @@ -4837,15 +4744,13 @@ [] l" by (import rich_list SUFFIX_DEF) -constdefs - UNZIP_FST :: "('a * 'b) list => 'a list" +definition UNZIP_FST :: "('a * 'b) list => 'a list" where "UNZIP_FST == %l::('a::type * 'b::type) list. fst (unzip l)" lemma UNZIP_FST_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_FST l = fst (unzip l)" by (import rich_list UNZIP_FST_DEF) -constdefs - UNZIP_SND :: "('a * 'b) list => 'b list" +definition UNZIP_SND :: "('a * 'b) list => 'b list" where "UNZIP_SND == %l::('a::type * 'b::type) list. snd (unzip l)" lemma UNZIP_SND_DEF: "ALL l::('a::type * 'b::type) list. UNZIP_SND l = snd (unzip l)" @@ -5916,8 +5821,7 @@ ;setup_theory state_transformer -constdefs - UNIT :: "'b => 'a => 'b * 'a" +definition UNIT :: "'b => 'a => 'b * 'a" where "(op ==::('b::type => 'a::type => 'b::type * 'a::type) => ('b::type => 'a::type => 'b::type * 'a::type) => prop) (UNIT::'b::type => 'a::type => 'b::type * 'a::type) @@ -5926,8 +5830,7 @@ lemma UNIT_DEF: "ALL x::'b::type. UNIT x = Pair x" by (import state_transformer UNIT_DEF) -constdefs - BIND :: "('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a" +definition BIND :: "('a => 'b * 'a) => ('b => 'a => 'c * 'a) => 'a => 'c * 'a" where "(op ==::(('a::type => 'b::type * 'a::type) => ('b::type => 'a::type => 'c::type * 'a::type) => 'a::type => 'c::type * 'a::type) @@ -5967,8 +5870,7 @@ g)))" by (import state_transformer BIND_DEF) -constdefs - MMAP :: "('c => 'b) => ('a => 'c * 'a) => 'a => 'b * 'a" +definition MMAP :: "('c => 'b) => ('a => 'c * 'a) => 'a => 'b * 'a" where "MMAP == %(f::'c::type => 'b::type) m::'a::type => 'c::type * 'a::type. BIND m (UNIT o f)" @@ -5977,8 +5879,7 @@ MMAP f m = BIND m (UNIT o f)" by (import state_transformer MMAP_DEF) -constdefs - JOIN :: "('a => ('a => 'b * 'a) * 'a) => 'a => 'b * 'a" +definition JOIN :: "('a => ('a => 'b * 'a) * 'a) => 'a => 'b * 'a" where "JOIN == %z::'a::type => ('a::type => 'b::type * 'a::type) * 'a::type. BIND z I" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Import/HOL/HOL4Prob.thy --- a/src/HOL/Import/HOL/HOL4Prob.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Import/HOL/HOL4Prob.thy Mon Mar 01 13:42:31 2010 +0100 @@ -373,8 +373,7 @@ alg_twin x y = (EX l::bool list. x = SNOC True l & y = SNOC False l)" by (import prob_canon alg_twin_def) -constdefs - alg_order_tupled :: "bool list * bool list => bool" +definition alg_order_tupled :: "bool list * bool list => bool" where "(op ==::(bool list * bool list => bool) => (bool list * bool list => bool) => prop) (alg_order_tupled::bool list * bool list => bool) @@ -1917,8 +1916,7 @@ P 0 & (ALL v::nat. P (Suc v div 2) --> P (Suc v)) --> All P" by (import prob_uniform unif_bound_ind) -constdefs - unif_tupled :: "nat * (nat => bool) => nat * (nat => bool)" +definition unif_tupled :: "nat * (nat => bool) => nat * (nat => bool)" where "unif_tupled == WFREC (SOME R::nat * (nat => bool) => nat * (nat => bool) => bool. @@ -1963,8 +1961,7 @@ (ALL v::nat. All (P v))" by (import prob_uniform unif_ind) -constdefs - uniform_tupled :: "nat * nat * (nat => bool) => nat * (nat => bool)" +definition uniform_tupled :: "nat * nat * (nat => bool) => nat * (nat => bool)" where "uniform_tupled == WFREC (SOME R::nat * nat * (nat => bool) => nat * nat * (nat => bool) => bool. diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Import/HOL/HOL4Real.thy --- a/src/HOL/Import/HOL/HOL4Real.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Import/HOL/HOL4Real.thy Mon Mar 01 13:42:31 2010 +0100 @@ -39,29 +39,25 @@ hreal_lt (hreal_add x y) (hreal_add x z) = hreal_lt y z" by (import realax HREAL_LT_LADD) -constdefs - treal_0 :: "hreal * hreal" +definition treal_0 :: "hreal * hreal" where "treal_0 == (hreal_1, hreal_1)" lemma treal_0: "treal_0 = (hreal_1, hreal_1)" by (import realax treal_0) -constdefs - treal_1 :: "hreal * hreal" +definition treal_1 :: "hreal * hreal" where "treal_1 == (hreal_add hreal_1 hreal_1, hreal_1)" lemma treal_1: "treal_1 = (hreal_add hreal_1 hreal_1, hreal_1)" by (import realax treal_1) -constdefs - treal_neg :: "hreal * hreal => hreal * hreal" +definition treal_neg :: "hreal * hreal => hreal * hreal" where "treal_neg == %(x::hreal, y::hreal). (y, x)" lemma treal_neg: "ALL (x::hreal) y::hreal. treal_neg (x, y) = (y, x)" by (import realax treal_neg) -constdefs - treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" +definition treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" where "treal_add == %(x1::hreal, y1::hreal) (x2::hreal, y2::hreal). (hreal_add x1 x2, hreal_add y1 y2)" @@ -70,8 +66,7 @@ treal_add (x1, y1) (x2, y2) = (hreal_add x1 x2, hreal_add y1 y2)" by (import realax treal_add) -constdefs - treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" +definition treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" where "treal_mul == %(x1::hreal, y1::hreal) (x2::hreal, y2::hreal). (hreal_add (hreal_mul x1 x2) (hreal_mul y1 y2), @@ -83,8 +78,7 @@ hreal_add (hreal_mul x1 y2) (hreal_mul y1 x2))" by (import realax treal_mul) -constdefs - treal_lt :: "hreal * hreal => hreal * hreal => bool" +definition treal_lt :: "hreal * hreal => hreal * hreal => bool" where "treal_lt == %(x1::hreal, y1::hreal) (x2::hreal, y2::hreal). hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)" @@ -93,8 +87,7 @@ treal_lt (x1, y1) (x2, y2) = hreal_lt (hreal_add x1 y2) (hreal_add x2 y1)" by (import realax treal_lt) -constdefs - treal_inv :: "hreal * hreal => hreal * hreal" +definition treal_inv :: "hreal * hreal => hreal * hreal" where "treal_inv == %(x::hreal, y::hreal). if x = y then treal_0 @@ -110,8 +103,7 @@ else (hreal_1, hreal_add (hreal_inv (hreal_sub y x)) hreal_1))" by (import realax treal_inv) -constdefs - treal_eq :: "hreal * hreal => hreal * hreal => bool" +definition treal_eq :: "hreal * hreal => hreal * hreal => bool" where "treal_eq == %(x1::hreal, y1::hreal) (x2::hreal, y2::hreal). hreal_add x1 y2 = hreal_add x2 y1" @@ -194,15 +186,13 @@ treal_lt treal_0 (treal_mul x y)" by (import realax TREAL_LT_MUL) -constdefs - treal_of_hreal :: "hreal => hreal * hreal" +definition treal_of_hreal :: "hreal => hreal * hreal" where "treal_of_hreal == %x::hreal. (hreal_add x hreal_1, hreal_1)" lemma treal_of_hreal: "ALL x::hreal. treal_of_hreal x = (hreal_add x hreal_1, hreal_1)" by (import realax treal_of_hreal) -constdefs - hreal_of_treal :: "hreal * hreal => hreal" +definition hreal_of_treal :: "hreal * hreal => hreal" where "hreal_of_treal == %(x::hreal, y::hreal). SOME d::hreal. x = hreal_add y d" lemma hreal_of_treal: "ALL (x::hreal) y::hreal. @@ -579,8 +569,7 @@ (EX x::real. ALL y::real. (EX x::real. P x & y < x) = (y < x))" by (import real REAL_SUP_EXISTS) -constdefs - sup :: "(real => bool) => real" +definition sup :: "(real => bool) => real" where "sup == %P::real => bool. SOME s::real. ALL y::real. (EX x::real. P x & y < x) = (y < s)" @@ -781,8 +770,7 @@ ;setup_theory topology -constdefs - re_Union :: "(('a => bool) => bool) => 'a => bool" +definition re_Union :: "(('a => bool) => bool) => 'a => bool" where "re_Union == %(P::('a::type => bool) => bool) x::'a::type. EX s::'a::type => bool. P s & s x" @@ -791,8 +779,7 @@ re_Union P = (%x::'a::type. EX s::'a::type => bool. P s & s x)" by (import topology re_Union) -constdefs - re_union :: "('a => bool) => ('a => bool) => 'a => bool" +definition re_union :: "('a => bool) => ('a => bool) => 'a => bool" where "re_union == %(P::'a::type => bool) (Q::'a::type => bool) x::'a::type. P x | Q x" @@ -800,8 +787,7 @@ re_union P Q = (%x::'a::type. P x | Q x)" by (import topology re_union) -constdefs - re_intersect :: "('a => bool) => ('a => bool) => 'a => bool" +definition re_intersect :: "('a => bool) => ('a => bool) => 'a => bool" where "re_intersect == %(P::'a::type => bool) (Q::'a::type => bool) x::'a::type. P x & Q x" @@ -809,22 +795,19 @@ re_intersect P Q = (%x::'a::type. P x & Q x)" by (import topology re_intersect) -constdefs - re_null :: "'a => bool" +definition re_null :: "'a => bool" where "re_null == %x::'a::type. False" lemma re_null: "re_null = (%x::'a::type. False)" by (import topology re_null) -constdefs - re_universe :: "'a => bool" +definition re_universe :: "'a => bool" where "re_universe == %x::'a::type. True" lemma re_universe: "re_universe = (%x::'a::type. True)" by (import topology re_universe) -constdefs - re_subset :: "('a => bool) => ('a => bool) => bool" +definition re_subset :: "('a => bool) => ('a => bool) => bool" where "re_subset == %(P::'a::type => bool) Q::'a::type => bool. ALL x::'a::type. P x --> Q x" @@ -832,8 +815,7 @@ re_subset P Q = (ALL x::'a::type. P x --> Q x)" by (import topology re_subset) -constdefs - re_compl :: "('a => bool) => 'a => bool" +definition re_compl :: "('a => bool) => 'a => bool" where "re_compl == %(P::'a::type => bool) x::'a::type. ~ P x" lemma re_compl: "ALL P::'a::type => bool. re_compl P = (%x::'a::type. ~ P x)" @@ -853,8 +835,7 @@ re_subset P Q & re_subset Q R --> re_subset P R" by (import topology SUBSET_TRANS) -constdefs - istopology :: "(('a => bool) => bool) => bool" +definition istopology :: "(('a => bool) => bool) => bool" where "istopology == %L::('a::type => bool) => bool. L re_null & @@ -900,8 +881,7 @@ re_subset xa (open x) --> open x (re_Union xa)" by (import topology TOPOLOGY_UNION) -constdefs - neigh :: "'a topology => ('a => bool) * 'a => bool" +definition neigh :: "'a topology => ('a => bool) * 'a => bool" where "neigh == %(top::'a::type topology) (N::'a::type => bool, x::'a::type). EX P::'a::type => bool. open top P & re_subset P N & P x" @@ -932,16 +912,14 @@ S' x --> (EX N::'a::type => bool. neigh top (N, x) & re_subset N S'))" by (import topology OPEN_NEIGH) -constdefs - closed :: "'a topology => ('a => bool) => bool" +definition closed :: "'a topology => ('a => bool) => bool" where "closed == %(L::'a::type topology) S'::'a::type => bool. open L (re_compl S')" lemma closed: "ALL (L::'a::type topology) S'::'a::type => bool. closed L S' = open L (re_compl S')" by (import topology closed) -constdefs - limpt :: "'a topology => 'a => ('a => bool) => bool" +definition limpt :: "'a topology => 'a => ('a => bool) => bool" where "limpt == %(top::'a::type topology) (x::'a::type) S'::'a::type => bool. ALL N::'a::type => bool. @@ -957,8 +935,7 @@ closed top S' = (ALL x::'a::type. limpt top x S' --> S' x)" by (import topology CLOSED_LIMPT) -constdefs - ismet :: "('a * 'a => real) => bool" +definition ismet :: "('a * 'a => real) => bool" where "ismet == %m::'a::type * 'a::type => real. (ALL (x::'a::type) y::'a::type. (m (x, y) = 0) = (x = y)) & @@ -1012,8 +989,7 @@ x ~= y --> 0 < dist m (x, y)" by (import topology METRIC_NZ) -constdefs - mtop :: "'a metric => 'a topology" +definition mtop :: "'a metric => 'a topology" where "mtop == %m::'a::type metric. topology @@ -1042,8 +1018,7 @@ S' xa --> (EX e>0. ALL y::'a::type. dist x (xa, y) < e --> S' y))" by (import topology MTOP_OPEN) -constdefs - B :: "'a metric => 'a * real => 'a => bool" +definition B :: "'a metric => 'a * real => 'a => bool" where "B == %(m::'a::type metric) (x::'a::type, e::real) y::'a::type. dist m (x, y) < e" @@ -1067,8 +1042,7 @@ lemma ISMET_R1: "ismet (%(x::real, y::real). abs (y - x))" by (import topology ISMET_R1) -constdefs - mr1 :: "real metric" +definition mr1 :: "real metric" where "mr1 == metric (%(x::real, y::real). abs (y - x))" lemma mr1: "mr1 = metric (%(x::real, y::real). abs (y - x))" @@ -1105,8 +1079,7 @@ ;setup_theory nets -constdefs - dorder :: "('a => 'a => bool) => bool" +definition dorder :: "('a => 'a => bool) => bool" where "dorder == %g::'a::type => 'a::type => bool. ALL (x::'a::type) y::'a::type. @@ -1120,8 +1093,7 @@ (EX z::'a::type. g z z & (ALL w::'a::type. g w z --> g w x & g w y)))" by (import nets dorder) -constdefs - tends :: "('b => 'a) => 'a => 'a topology * ('b => 'b => bool) => bool" +definition tends :: "('b => 'a) => 'a => 'a topology * ('b => 'b => bool) => bool" where "tends == %(s::'b::type => 'a::type) (l::'a::type) (top::'a::type topology, g::'b::type => 'b::type => bool). @@ -1137,8 +1109,7 @@ (EX n::'b::type. g n n & (ALL m::'b::type. g m n --> N (s m))))" by (import nets tends) -constdefs - bounded :: "'a metric * ('b => 'b => bool) => ('b => 'a) => bool" +definition bounded :: "'a metric * ('b => 'b => bool) => ('b => 'a) => bool" where "bounded == %(m::'a::type metric, g::'b::type => 'b::type => bool) f::'b::type => 'a::type. @@ -1152,8 +1123,7 @@ g N N & (ALL n::'b::type. g n N --> dist m (f n, x) < k))" by (import nets bounded) -constdefs - tendsto :: "'a metric * 'a => 'a => 'a => bool" +definition tendsto :: "'a metric * 'a => 'a => 'a => bool" where "tendsto == %(m::'a::type metric, x::'a::type) (y::'a::type) z::'a::type. 0 < dist m (x, y) & dist m (x, y) <= dist m (x, z)" @@ -1366,15 +1336,13 @@ hol4--> x x1 & hol4--> x x2 --> x1 = x2" by (import seq SEQ_UNIQ) -constdefs - convergent :: "(nat => real) => bool" +definition convergent :: "(nat => real) => bool" where "convergent == %f::nat => real. Ex (hol4--> f)" lemma convergent: "ALL f::nat => real. convergent f = Ex (hol4--> f)" by (import seq convergent) -constdefs - cauchy :: "(nat => real) => bool" +definition cauchy :: "(nat => real) => bool" where "cauchy == %f::nat => real. ALL e>0. @@ -1388,8 +1356,7 @@ ALL (m::nat) n::nat. N <= m & N <= n --> abs (f m - f n) < e)" by (import seq cauchy) -constdefs - lim :: "(nat => real) => real" +definition lim :: "(nat => real) => real" where "lim == %f::nat => real. Eps (hol4--> f)" lemma lim: "ALL f::nat => real. lim f = Eps (hol4--> f)" @@ -1398,8 +1365,7 @@ lemma SEQ_LIM: "ALL f::nat => real. convergent f = hol4--> f (lim f)" by (import seq SEQ_LIM) -constdefs - subseq :: "(nat => nat) => bool" +definition subseq :: "(nat => nat) => bool" where "subseq == %f::nat => nat. ALL (m::nat) n::nat. m < n --> f m < f n" lemma subseq: "ALL f::nat => nat. subseq f = (ALL (m::nat) n::nat. m < n --> f m < f n)" @@ -1541,23 +1507,20 @@ (ALL (a::real) b::real. a <= b --> P (a, b))" by (import seq BOLZANO_LEMMA) -constdefs - sums :: "(nat => real) => real => bool" +definition sums :: "(nat => real) => real => bool" where "sums == %f::nat => real. hol4--> (%n::nat. real.sum (0, n) f)" lemma sums: "ALL (f::nat => real) s::real. sums f s = hol4--> (%n::nat. real.sum (0, n) f) s" by (import seq sums) -constdefs - summable :: "(nat => real) => bool" +definition summable :: "(nat => real) => bool" where "summable == %f::nat => real. Ex (sums f)" lemma summable: "ALL f::nat => real. summable f = Ex (sums f)" by (import seq summable) -constdefs - suminf :: "(nat => real) => real" +definition suminf :: "(nat => real) => real" where "suminf == %f::nat => real. Eps (sums f)" lemma suminf: "ALL f::nat => real. suminf f = Eps (sums f)" @@ -1692,8 +1655,7 @@ ;setup_theory lim -constdefs - tends_real_real :: "(real => real) => real => real => bool" +definition tends_real_real :: "(real => real) => real => real => bool" where "tends_real_real == %(f::real => real) (l::real) x0::real. tends f l (mtop mr1, tendsto (mr1, x0))" @@ -1763,8 +1725,7 @@ tends_real_real f l x0" by (import lim LIM_TRANSFORM) -constdefs - diffl :: "(real => real) => real => real => bool" +definition diffl :: "(real => real) => real => real => bool" where "diffl == %(f::real => real) (l::real) x::real. tends_real_real (%h::real. (f (x + h) - f x) / h) l 0" @@ -1773,8 +1734,7 @@ diffl f l x = tends_real_real (%h::real. (f (x + h) - f x) / h) l 0" by (import lim diffl) -constdefs - contl :: "(real => real) => real => bool" +definition contl :: "(real => real) => real => bool" where "contl == %(f::real => real) x::real. tends_real_real (%h::real. f (x + h)) (f x) 0" @@ -1782,8 +1742,7 @@ contl f x = tends_real_real (%h::real. f (x + h)) (f x) 0" by (import lim contl) -constdefs - differentiable :: "(real => real) => real => bool" +definition differentiable :: "(real => real) => real => bool" where "differentiable == %(f::real => real) x::real. EX l::real. diffl f l x" lemma differentiable: "ALL (f::real => real) x::real. @@ -2127,8 +2086,7 @@ summable (%n::nat. f n * z ^ n)" by (import powser POWSER_INSIDE) -constdefs - diffs :: "(nat => real) => nat => real" +definition diffs :: "(nat => real) => nat => real" where "diffs == %(c::nat => real) n::nat. real (Suc n) * c (Suc n)" lemma diffs: "ALL c::nat => real. diffs c = (%n::nat. real (Suc n) * c (Suc n))" @@ -2204,15 +2162,13 @@ ;setup_theory transc -constdefs - exp :: "real => real" +definition exp :: "real => real" where "exp == %x::real. suminf (%n::nat. inverse (real (FACT n)) * x ^ n)" lemma exp: "ALL x::real. exp x = suminf (%n::nat. inverse (real (FACT n)) * x ^ n)" by (import transc exp) -constdefs - cos :: "real => real" +definition cos :: "real => real" where "cos == %x::real. suminf @@ -2226,8 +2182,7 @@ (if EVEN n then (- 1) ^ (n div 2) / real (FACT n) else 0) * x ^ n)" by (import transc cos) -constdefs - sin :: "real => real" +definition sin :: "real => real" where "sin == %x::real. suminf @@ -2364,8 +2319,7 @@ lemma EXP_TOTAL: "ALL y>0. EX x::real. exp x = y" by (import transc EXP_TOTAL) -constdefs - ln :: "real => real" +definition ln :: "real => real" where "ln == %x::real. SOME u::real. exp u = x" lemma ln: "ALL x::real. ln x = (SOME u::real. exp u = x)" @@ -2410,16 +2364,14 @@ lemma LN_POS: "ALL x>=1. 0 <= ln x" by (import transc LN_POS) -constdefs - root :: "nat => real => real" +definition root :: "nat => real => real" where "root == %(n::nat) x::real. SOME u::real. (0 < x --> 0 < u) & u ^ n = x" lemma root: "ALL (n::nat) x::real. root n x = (SOME u::real. (0 < x --> 0 < u) & u ^ n = x)" by (import transc root) -constdefs - sqrt :: "real => real" +definition sqrt :: "real => real" where "sqrt == root 2" lemma sqrt: "ALL x::real. sqrt x = root 2 x" @@ -2584,8 +2536,7 @@ lemma COS_ISZERO: "EX! x::real. 0 <= x & x <= 2 & cos x = 0" by (import transc COS_ISZERO) -constdefs - pi :: "real" +definition pi :: "real" where "pi == 2 * (SOME x::real. 0 <= x & x <= 2 & cos x = 0)" lemma pi: "pi = 2 * (SOME x::real. 0 <= x & x <= 2 & cos x = 0)" @@ -2689,8 +2640,7 @@ (EX n::nat. EVEN n & x = - (real n * (pi / 2))))" by (import transc SIN_ZERO) -constdefs - tan :: "real => real" +definition tan :: "real => real" where "tan == %x::real. sin x / cos x" lemma tan: "ALL x::real. tan x = sin x / cos x" @@ -2736,23 +2686,20 @@ lemma TAN_TOTAL: "ALL y::real. EX! x::real. - (pi / 2) < x & x < pi / 2 & tan x = y" by (import transc TAN_TOTAL) -constdefs - asn :: "real => real" +definition asn :: "real => real" where "asn == %y::real. SOME x::real. - (pi / 2) <= x & x <= pi / 2 & sin x = y" lemma asn: "ALL y::real. asn y = (SOME x::real. - (pi / 2) <= x & x <= pi / 2 & sin x = y)" by (import transc asn) -constdefs - acs :: "real => real" +definition acs :: "real => real" where "acs == %y::real. SOME x::real. 0 <= x & x <= pi & cos x = y" lemma acs: "ALL y::real. acs y = (SOME x::real. 0 <= x & x <= pi & cos x = y)" by (import transc acs) -constdefs - atn :: "real => real" +definition atn :: "real => real" where "atn == %y::real. SOME x::real. - (pi / 2) < x & x < pi / 2 & tan x = y" lemma atn: "ALL y::real. atn y = (SOME x::real. - (pi / 2) < x & x < pi / 2 & tan x = y)" @@ -2845,8 +2792,7 @@ lemma DIFF_ATN: "ALL x::real. diffl atn (inverse (1 + x ^ 2)) x" by (import transc DIFF_ATN) -constdefs - division :: "real * real => (nat => real) => bool" +definition division :: "real * real => (nat => real) => bool" where "(op ==::(real * real => (nat => real) => bool) => (real * real => (nat => real) => bool) => prop) (division::real * real => (nat => real) => bool) @@ -2898,8 +2844,7 @@ b)))))))))" by (import transc division) -constdefs - dsize :: "(nat => real) => nat" +definition dsize :: "(nat => real) => nat" where "(op ==::((nat => real) => nat) => ((nat => real) => nat) => prop) (dsize::(nat => real) => nat) (%D::nat => real. @@ -2937,8 +2882,7 @@ ((op =::real => real => bool) (D n) (D N)))))))" by (import transc dsize) -constdefs - tdiv :: "real * real => (nat => real) * (nat => real) => bool" +definition tdiv :: "real * real => (nat => real) * (nat => real) => bool" where "tdiv == %(a::real, b::real) (D::nat => real, p::nat => real). division (a, b) D & (ALL n::nat. D n <= p n & p n <= D (Suc n))" @@ -2948,16 +2892,14 @@ (division (a, b) D & (ALL n::nat. D n <= p n & p n <= D (Suc n)))" by (import transc tdiv) -constdefs - gauge :: "(real => bool) => (real => real) => bool" +definition gauge :: "(real => bool) => (real => real) => bool" where "gauge == %(E::real => bool) g::real => real. ALL x::real. E x --> 0 < g x" lemma gauge: "ALL (E::real => bool) g::real => real. gauge E g = (ALL x::real. E x --> 0 < g x)" by (import transc gauge) -constdefs - fine :: "(real => real) => (nat => real) * (nat => real) => bool" +definition fine :: "(real => real) => (nat => real) * (nat => real) => bool" where "(op ==::((real => real) => (nat => real) * (nat => real) => bool) => ((real => real) => (nat => real) * (nat => real) => bool) => prop) @@ -3000,8 +2942,7 @@ (g (p n))))))))" by (import transc fine) -constdefs - rsum :: "(nat => real) * (nat => real) => (real => real) => real" +definition rsum :: "(nat => real) * (nat => real) => (real => real) => real" where "rsum == %(D::nat => real, p::nat => real) f::real => real. real.sum (0, dsize D) (%n::nat. f (p n) * (D (Suc n) - D n))" @@ -3011,8 +2952,7 @@ real.sum (0, dsize D) (%n::nat. f (p n) * (D (Suc n) - D n))" by (import transc rsum) -constdefs - Dint :: "real * real => (real => real) => real => bool" +definition Dint :: "real * real => (real => real) => real => bool" where "Dint == %(a::real, b::real) (f::real => real) k::real. ALL e>0. @@ -3313,8 +3253,7 @@ poly_diff_aux n (h # t) = real n * h # poly_diff_aux (Suc n) t)" by (import poly poly_diff_aux_def) -constdefs - diff :: "real list => real list" +definition diff :: "real list => real list" where "diff == %l::real list. if l = [] then [] else poly_diff_aux 1 (tl l)" lemma poly_diff_def: "ALL l::real list. diff l = (if l = [] then [] else poly_diff_aux 1 (tl l))" @@ -3622,8 +3561,7 @@ poly p = poly q --> poly (diff p) = poly (diff q)" by (import poly POLY_DIFF_WELLDEF) -constdefs - poly_divides :: "real list => real list => bool" +definition poly_divides :: "real list => real list => bool" where "poly_divides == %(p1::real list) p2::real list. EX q::real list. poly p2 = poly (poly_mul p1 q)" @@ -3681,8 +3619,7 @@ ~ poly_divides (poly_exp [- a, 1] (Suc n)) p)" by (import poly POLY_ORDER) -constdefs - poly_order :: "real => real list => nat" +definition poly_order :: "real => real list => nat" where "poly_order == %(a::real) p::real list. SOME n::nat. @@ -3754,8 +3691,7 @@ (ALL a::real. poly_order a q = (if poly_order a p = 0 then 0 else 1))" by (import poly POLY_SQUAREFREE_DECOMP_ORDER) -constdefs - rsquarefree :: "real list => bool" +definition rsquarefree :: "real list => bool" where "rsquarefree == %p::real list. poly p ~= poly [] & @@ -3798,8 +3734,7 @@ lemma POLY_NORMALIZE: "ALL t::real list. poly (normalize t) = poly t" by (import poly POLY_NORMALIZE) -constdefs - degree :: "real list => nat" +definition degree :: "real list => nat" where "degree == %p::real list. PRE (length (normalize p))" lemma degree: "ALL p::real list. degree p = PRE (length (normalize p))" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Import/HOL/HOL4Vec.thy --- a/src/HOL/Import/HOL/HOL4Vec.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Import/HOL/HOL4Vec.thy Mon Mar 01 13:42:31 2010 +0100 @@ -164,8 +164,7 @@ lemma word_base0_def: "word_base0 = (%a::'a::type list. mk_word (CONSTR 0 a (%n::nat. BOTTOM)))" by (import word_base word_base0_def) -constdefs - WORD :: "'a list => 'a word" +definition WORD :: "'a list => 'a word" where "WORD == word_base0" lemma WORD: "WORD = word_base0" @@ -680,8 +679,7 @@ ;setup_theory word_num -constdefs - LVAL :: "('a => nat) => nat => 'a list => nat" +definition LVAL :: "('a => nat) => nat => 'a list => nat" where "LVAL == %(f::'a::type => nat) b::nat. foldl (%(e::nat) x::'a::type. b * e + f x) 0" @@ -756,8 +754,7 @@ SNOC (frep (m mod b)) (NLIST n frep b (m div b)))" by (import word_num NLIST_DEF) -constdefs - NWORD :: "nat => (nat => 'a) => nat => nat => 'a word" +definition NWORD :: "nat => (nat => 'a) => nat => nat => 'a word" where "NWORD == %(n::nat) (frep::nat => 'a::type) (b::nat) m::nat. WORD (NLIST n frep b m)" @@ -1076,8 +1073,7 @@ EXISTSABIT P (WCAT (w1, w2)) = (EXISTSABIT P w1 | EXISTSABIT P w2)" by (import word_bitop EXISTSABIT_WCAT) -constdefs - SHR :: "bool => 'a => 'a word => 'a word * 'a" +definition SHR :: "bool => 'a => 'a word => 'a word * 'a" where "SHR == %(f::bool) (b::'a::type) w::'a::type word. (WCAT @@ -1093,8 +1089,7 @@ bit 0 w)" by (import word_bitop SHR_DEF) -constdefs - SHL :: "bool => 'a word => 'a => 'a * 'a word" +definition SHL :: "bool => 'a word => 'a => 'a * 'a word" where "SHL == %(f::bool) (w::'a::type word) b::'a::type. (bit (PRE (WORDLEN w)) w, @@ -1196,8 +1191,7 @@ ;setup_theory bword_num -constdefs - BV :: "bool => nat" +definition BV :: "bool => nat" where "BV == %b::bool. if b then Suc 0 else 0" lemma BV_DEF: "ALL b::bool. BV b = (if b then Suc 0 else 0)" @@ -1248,15 +1242,13 @@ BNVAL (WCAT (w1, w2)) = BNVAL w1 * 2 ^ m + BNVAL w2))" by (import bword_num BNVAL_WCAT) -constdefs - VB :: "nat => bool" +definition VB :: "nat => bool" where "VB == %n::nat. n mod 2 ~= 0" lemma VB_DEF: "ALL n::nat. VB n = (n mod 2 ~= 0)" by (import bword_num VB_DEF) -constdefs - NBWORD :: "nat => nat => bool word" +definition NBWORD :: "nat => nat => bool word" where "NBWORD == %(n::nat) m::nat. WORD (NLIST n VB 2 m)" lemma NBWORD_DEF: "ALL (n::nat) m::nat. NBWORD n m = WORD (NLIST n VB 2 m)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Import/HOL/HOL4Word32.thy --- a/src/HOL/Import/HOL/HOL4Word32.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Import/HOL/HOL4Word32.thy Mon Mar 01 13:42:31 2010 +0100 @@ -68,8 +68,7 @@ BITS h l n = MOD_2EXP (Suc h - l) (DIV_2EXP l n)" by (import bits BITS_def) -constdefs - bit :: "nat => nat => bool" +definition bit :: "nat => nat => bool" where "bit == %(b::nat) n::nat. BITS b b n = 1" lemma BIT_def: "ALL (b::nat) n::nat. bit b n = (BITS b b n = 1)" @@ -840,15 +839,13 @@ lemma w_T_def: "w_T = mk_word32 (EQUIV COMP0)" by (import word32 w_T_def) -constdefs - word_suc :: "word32 => word32" +definition word_suc :: "word32 => word32" where "word_suc == %T1::word32. mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))" lemma word_suc: "ALL T1::word32. word_suc T1 = mk_word32 (EQUIV (Suc (Eps (dest_word32 T1))))" by (import word32 word_suc) -constdefs - word_add :: "word32 => word32 => word32" +definition word_add :: "word32 => word32 => word32" where "word_add == %(T1::word32) T2::word32. mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))" @@ -858,8 +855,7 @@ mk_word32 (EQUIV (Eps (dest_word32 T1) + Eps (dest_word32 T2)))" by (import word32 word_add) -constdefs - word_mul :: "word32 => word32 => word32" +definition word_mul :: "word32 => word32 => word32" where "word_mul == %(T1::word32) T2::word32. mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))" @@ -869,8 +865,7 @@ mk_word32 (EQUIV (Eps (dest_word32 T1) * Eps (dest_word32 T2)))" by (import word32 word_mul) -constdefs - word_1comp :: "word32 => word32" +definition word_1comp :: "word32 => word32" where "word_1comp == %T1::word32. mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))" @@ -878,8 +873,7 @@ word_1comp T1 = mk_word32 (EQUIV (ONE_COMP (Eps (dest_word32 T1))))" by (import word32 word_1comp) -constdefs - word_2comp :: "word32 => word32" +definition word_2comp :: "word32 => word32" where "word_2comp == %T1::word32. mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))" @@ -887,24 +881,21 @@ word_2comp T1 = mk_word32 (EQUIV (TWO_COMP (Eps (dest_word32 T1))))" by (import word32 word_2comp) -constdefs - word_lsr1 :: "word32 => word32" +definition word_lsr1 :: "word32 => word32" where "word_lsr1 == %T1::word32. mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))" lemma word_lsr1: "ALL T1::word32. word_lsr1 T1 = mk_word32 (EQUIV (LSR_ONE (Eps (dest_word32 T1))))" by (import word32 word_lsr1) -constdefs - word_asr1 :: "word32 => word32" +definition word_asr1 :: "word32 => word32" where "word_asr1 == %T1::word32. mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))" lemma word_asr1: "ALL T1::word32. word_asr1 T1 = mk_word32 (EQUIV (ASR_ONE (Eps (dest_word32 T1))))" by (import word32 word_asr1) -constdefs - word_ror1 :: "word32 => word32" +definition word_ror1 :: "word32 => word32" where "word_ror1 == %T1::word32. mk_word32 (EQUIV (ROR_ONE (Eps (dest_word32 T1))))" lemma word_ror1: "ALL T1::word32. @@ -940,8 +931,7 @@ lemma MSB_def: "ALL T1::word32. MSB T1 = MSBn (Eps (dest_word32 T1))" by (import word32 MSB_def) -constdefs - bitwise_or :: "word32 => word32 => word32" +definition bitwise_or :: "word32 => word32 => word32" where "bitwise_or == %(T1::word32) T2::word32. mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" @@ -951,8 +941,7 @@ mk_word32 (EQUIV (OR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" by (import word32 bitwise_or) -constdefs - bitwise_eor :: "word32 => word32 => word32" +definition bitwise_eor :: "word32 => word32 => word32" where "bitwise_eor == %(T1::word32) T2::word32. mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" @@ -962,8 +951,7 @@ mk_word32 (EQUIV (EOR (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" by (import word32 bitwise_eor) -constdefs - bitwise_and :: "word32 => word32 => word32" +definition bitwise_and :: "word32 => word32 => word32" where "bitwise_and == %(T1::word32) T2::word32. mk_word32 (EQUIV (AND (Eps (dest_word32 T1)) (Eps (dest_word32 T2))))" @@ -1154,36 +1142,31 @@ lemma ADD_TWO_COMP2: "ALL x::word32. word_add (word_2comp x) x = w_0" by (import word32 ADD_TWO_COMP2) -constdefs - word_sub :: "word32 => word32 => word32" +definition word_sub :: "word32 => word32 => word32" where "word_sub == %(a::word32) b::word32. word_add a (word_2comp b)" lemma word_sub: "ALL (a::word32) b::word32. word_sub a b = word_add a (word_2comp b)" by (import word32 word_sub) -constdefs - word_lsl :: "word32 => nat => word32" +definition word_lsl :: "word32 => nat => word32" where "word_lsl == %(a::word32) n::nat. word_mul a (n2w (2 ^ n))" lemma word_lsl: "ALL (a::word32) n::nat. word_lsl a n = word_mul a (n2w (2 ^ n))" by (import word32 word_lsl) -constdefs - word_lsr :: "word32 => nat => word32" +definition word_lsr :: "word32 => nat => word32" where "word_lsr == %(a::word32) n::nat. (word_lsr1 ^^ n) a" lemma word_lsr: "ALL (a::word32) n::nat. word_lsr a n = (word_lsr1 ^^ n) a" by (import word32 word_lsr) -constdefs - word_asr :: "word32 => nat => word32" +definition word_asr :: "word32 => nat => word32" where "word_asr == %(a::word32) n::nat. (word_asr1 ^^ n) a" lemma word_asr: "ALL (a::word32) n::nat. word_asr a n = (word_asr1 ^^ n) a" by (import word32 word_asr) -constdefs - word_ror :: "word32 => nat => word32" +definition word_ror :: "word32 => nat => word32" where "word_ror == %(a::word32) n::nat. (word_ror1 ^^ n) a" lemma word_ror: "ALL (a::word32) n::nat. word_ror a n = (word_ror1 ^^ n) a" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Import/HOL4Compat.thy --- a/src/HOL/Import/HOL4Compat.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Import/HOL4Compat.thy Mon Mar 01 13:42:31 2010 +0100 @@ -15,8 +15,7 @@ lemma COND_DEF:"(If b t f) = (@x. ((b = True) --> (x = t)) & ((b = False) --> (x = f)))" by auto -constdefs - LET :: "['a \ 'b,'a] \ 'b" +definition LET :: "['a \ 'b,'a] \ 'b" where "LET f s == f s" lemma [hol4rew]: "LET f s = Let s f" @@ -119,10 +118,10 @@ lemma LESS_OR_EQ: "m <= (n::nat) = (m < n | m = n)" by auto -constdefs - nat_gt :: "nat => nat => bool" +definition nat_gt :: "nat => nat => bool" where "nat_gt == %m n. n < m" - nat_ge :: "nat => nat => bool" + +definition nat_ge :: "nat => nat => bool" where "nat_ge == %m n. nat_gt m n | m = n" lemma [hol4rew]: "nat_gt m n = (n < m)" @@ -200,8 +199,7 @@ qed qed; -constdefs - FUNPOW :: "('a => 'a) => nat => 'a => 'a" +definition FUNPOW :: "('a => 'a) => nat => 'a => 'a" where "FUNPOW f n == f ^^ n" lemma FUNPOW: "(ALL f x. (f ^^ 0) x = x) & @@ -229,14 +227,16 @@ lemma DIVISION: "(0::nat) < n --> (!k. (k = k div n * n + k mod n) & k mod n < n)" by simp -constdefs - ALT_ZERO :: nat +definition ALT_ZERO :: nat where "ALT_ZERO == 0" - NUMERAL_BIT1 :: "nat \ nat" + +definition NUMERAL_BIT1 :: "nat \ nat" where "NUMERAL_BIT1 n == n + (n + Suc 0)" - NUMERAL_BIT2 :: "nat \ nat" + +definition NUMERAL_BIT2 :: "nat \ nat" where "NUMERAL_BIT2 n == n + (n + Suc (Suc 0))" - NUMERAL :: "nat \ nat" + +definition NUMERAL :: "nat \ nat" where "NUMERAL x == x" lemma [hol4rew]: "NUMERAL ALT_ZERO = 0" @@ -329,8 +329,7 @@ lemma NULL_DEF: "(null [] = True) & (!h t. null (h # t) = False)" by simp -constdefs - sum :: "nat list \ nat" +definition sum :: "nat list \ nat" where "sum l == foldr (op +) l 0" lemma SUM: "(sum [] = 0) & (!h t. sum (h#t) = h + sum t)" @@ -359,8 +358,7 @@ (ALL n x. replicate (Suc n) x = x # replicate n x)" by simp -constdefs - FOLDR :: "[['a,'b]\'b,'b,'a list] \ 'b" +definition FOLDR :: "[['a,'b]\'b,'b,'a list] \ 'b" where "FOLDR f e l == foldr f l e" lemma [hol4rew]: "FOLDR f e l = foldr f l e" @@ -418,8 +416,7 @@ lemma list_CASES: "(l = []) | (? t h. l = h#t)" by (induct l,auto) -constdefs - ZIP :: "'a list * 'b list \ ('a * 'b) list" +definition ZIP :: "'a list * 'b list \ ('a * 'b) list" where "ZIP == %(a,b). zip a b" lemma ZIP: "(zip [] [] = []) & @@ -514,8 +511,7 @@ lemma pow: "(!x::real. x ^ 0 = 1) & (!x::real. ALL n. x ^ (Suc n) = x * x ^ n)" by simp -constdefs - real_gt :: "real => real => bool" +definition real_gt :: "real => real => bool" where "real_gt == %x y. y < x" lemma [hol4rew]: "real_gt x y = (y < x)" @@ -524,8 +520,7 @@ lemma real_gt: "ALL x (y::real). (y < x) = (y < x)" by simp -constdefs - real_ge :: "real => real => bool" +definition real_ge :: "real => real => bool" where "real_ge x y == y <= x" lemma [hol4rew]: "real_ge x y = (y <= x)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Import/HOLLight/HOLLight.thy --- a/src/HOL/Import/HOLLight/HOLLight.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Import/HOLLight/HOLLight.thy Mon Mar 01 13:42:31 2010 +0100 @@ -95,8 +95,7 @@ lemma EXCLUDED_MIDDLE: "ALL t::bool. t | ~ t" by (import hollight EXCLUDED_MIDDLE) -constdefs - COND :: "bool => 'A => 'A => 'A" +definition COND :: "bool => 'A => 'A => 'A" where "COND == %(t::bool) (t1::'A::type) t2::'A::type. SOME x::'A::type. (t = True --> x = t1) & (t = False --> x = t2)" @@ -173,15 +172,13 @@ (b & P x True | ~ b & P y False)" by (import hollight th_cond) -constdefs - LET_END :: "'A => 'A" +definition LET_END :: "'A => 'A" where "LET_END == %t::'A::type. t" lemma DEF_LET_END: "LET_END = (%t::'A::type. t)" by (import hollight DEF_LET_END) -constdefs - GABS :: "('A => bool) => 'A" +definition GABS :: "('A => bool) => 'A" where "(op ==::(('A::type => bool) => 'A::type) => (('A::type => bool) => 'A::type) => prop) (GABS::('A::type => bool) => 'A::type) @@ -193,8 +190,7 @@ (Eps::('A::type => bool) => 'A::type)" by (import hollight DEF_GABS) -constdefs - GEQ :: "'A => 'A => bool" +definition GEQ :: "'A => 'A => bool" where "(op ==::('A::type => 'A::type => bool) => ('A::type => 'A::type => bool) => prop) (GEQ::'A::type => 'A::type => bool) (op =::'A::type => 'A::type => bool)" @@ -208,8 +204,7 @@ x = Pair_Rep a b" by (import hollight PAIR_EXISTS_THM) -constdefs - CURRY :: "('A * 'B => 'C) => 'A => 'B => 'C" +definition CURRY :: "('A * 'B => 'C) => 'A => 'B => 'C" where "CURRY == %(u::'A::type * 'B::type => 'C::type) (ua::'A::type) ub::'B::type. u (ua, ub)" @@ -219,8 +214,7 @@ u (ua, ub))" by (import hollight DEF_CURRY) -constdefs - UNCURRY :: "('A => 'B => 'C) => 'A * 'B => 'C" +definition UNCURRY :: "('A => 'B => 'C) => 'A * 'B => 'C" where "UNCURRY == %(u::'A::type => 'B::type => 'C::type) ua::'A::type * 'B::type. u (fst ua) (snd ua)" @@ -230,8 +224,7 @@ u (fst ua) (snd ua))" by (import hollight DEF_UNCURRY) -constdefs - PASSOC :: "(('A * 'B) * 'C => 'D) => 'A * 'B * 'C => 'D" +definition PASSOC :: "(('A * 'B) * 'C => 'D) => 'A * 'B * 'C => 'D" where "PASSOC == %(u::('A::type * 'B::type) * 'C::type => 'D::type) ua::'A::type * 'B::type * 'C::type. @@ -291,8 +284,7 @@ (m * n = NUMERAL_BIT1 0) = (m = NUMERAL_BIT1 0 & n = NUMERAL_BIT1 0)" by (import hollight MULT_EQ_1) -constdefs - EXP :: "nat => nat => nat" +definition EXP :: "nat => nat => nat" where "EXP == SOME EXP::nat => nat => nat. (ALL m::nat. EXP m 0 = NUMERAL_BIT1 0) & @@ -549,8 +541,7 @@ (EX m::nat. P m & (ALL x::nat. P x --> <= x m))" by (import hollight num_MAX) -constdefs - EVEN :: "nat => bool" +definition EVEN :: "nat => bool" where "EVEN == SOME EVEN::nat => bool. EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n))" @@ -560,8 +551,7 @@ EVEN 0 = True & (ALL n::nat. EVEN (Suc n) = (~ EVEN n)))" by (import hollight DEF_EVEN) -constdefs - ODD :: "nat => bool" +definition ODD :: "nat => bool" where "ODD == SOME ODD::nat => bool. ODD 0 = False & (ALL n::nat. ODD (Suc n) = (~ ODD n))" @@ -641,8 +631,7 @@ lemma SUC_SUB1: "ALL x::nat. Suc x - NUMERAL_BIT1 0 = x" by (import hollight SUC_SUB1) -constdefs - FACT :: "nat => nat" +definition FACT :: "nat => nat" where "FACT == SOME FACT::nat => nat. FACT 0 = NUMERAL_BIT1 0 & (ALL n::nat. FACT (Suc n) = Suc n * FACT n)" @@ -669,8 +658,7 @@ COND (n = 0) (x = 0 & xa = 0) (m = x * n + xa & < xa n)" by (import hollight DIVMOD_EXIST_0) -constdefs - DIV :: "nat => nat => nat" +definition DIV :: "nat => nat => nat" where "DIV == SOME q::nat => nat => nat. EX r::nat => nat => nat. @@ -686,8 +674,7 @@ (m = q m n * n + r m n & < (r m n) n))" by (import hollight DEF_DIV) -constdefs - MOD :: "nat => nat => nat" +definition MOD :: "nat => nat => nat" where "MOD == SOME r::nat => nat => nat. ALL (m::nat) n::nat. @@ -877,8 +864,7 @@ n ~= 0 & (ALL (q::nat) r::nat. m = q * n + r & < r n --> P q r))" by (import hollight DIVMOD_ELIM_THM) -constdefs - eqeq :: "'q_9910 => 'q_9909 => ('q_9910 => 'q_9909 => bool) => bool" +definition eqeq :: "'q_9910 => 'q_9909 => ('q_9910 => 'q_9909 => bool) => bool" where "eqeq == %(u::'q_9910::type) (ua::'q_9909::type) ub::'q_9910::type => 'q_9909::type => bool. ub u ua" @@ -888,8 +874,7 @@ ub::'q_9910::type => 'q_9909::type => bool. ub u ua)" by (import hollight DEF__equal__equal_) -constdefs - mod_nat :: "nat => nat => nat => bool" +definition mod_nat :: "nat => nat => nat => bool" where "mod_nat == %(u::nat) (ua::nat) ub::nat. EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2" @@ -898,8 +883,7 @@ EX (q1::nat) q2::nat. ua + u * q1 = ub + u * q2)" by (import hollight DEF_mod_nat) -constdefs - minimal :: "(nat => bool) => nat" +definition minimal :: "(nat => bool) => nat" where "minimal == %u::nat => bool. SOME n::nat. u n & (ALL m::nat. < m n --> ~ u m)" lemma DEF_minimal: "minimal = @@ -910,8 +894,7 @@ Ex P = (P (minimal P) & (ALL x::nat. < x (minimal P) --> ~ P x))" by (import hollight MINIMAL) -constdefs - WF :: "('A => 'A => bool) => bool" +definition WF :: "('A => 'A => bool) => bool" where "WF == %u::'A::type => 'A::type => bool. ALL P::'A::type => bool. @@ -1605,8 +1588,7 @@ ALL (xa::'A::type) y::'B::type. x (P xa y) = xa & Y (P xa y) = y)" by (import hollight INJ_INVERSE2) -constdefs - NUMPAIR :: "nat => nat => nat" +definition NUMPAIR :: "nat => nat => nat" where "NUMPAIR == %(u::nat) ua::nat. EXP (NUMERAL_BIT0 (NUMERAL_BIT1 0)) u * @@ -1626,8 +1608,7 @@ (NUMPAIR x1 y1 = NUMPAIR x2 y2) = (x1 = x2 & y1 = y2)" by (import hollight NUMPAIR_INJ) -constdefs - NUMFST :: "nat => nat" +definition NUMFST :: "nat => nat" where "NUMFST == SOME X::nat => nat. EX Y::nat => nat. @@ -1639,8 +1620,7 @@ ALL (x::nat) y::nat. X (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)" by (import hollight DEF_NUMFST) -constdefs - NUMSND :: "nat => nat" +definition NUMSND :: "nat => nat" where "NUMSND == SOME Y::nat => nat. ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y" @@ -1650,8 +1630,7 @@ ALL (x::nat) y::nat. NUMFST (NUMPAIR x y) = x & Y (NUMPAIR x y) = y)" by (import hollight DEF_NUMSND) -constdefs - NUMSUM :: "bool => nat => nat" +definition NUMSUM :: "bool => nat => nat" where "NUMSUM == %(u::bool) ua::nat. COND u (Suc (NUMERAL_BIT0 (NUMERAL_BIT1 0) * ua)) @@ -1667,8 +1646,7 @@ (NUMSUM b1 x1 = NUMSUM b2 x2) = (b1 = b2 & x1 = x2)" by (import hollight NUMSUM_INJ) -constdefs - NUMLEFT :: "nat => bool" +definition NUMLEFT :: "nat => bool" where "NUMLEFT == SOME X::nat => bool. EX Y::nat => nat. @@ -1680,8 +1658,7 @@ ALL (x::bool) y::nat. X (NUMSUM x y) = x & Y (NUMSUM x y) = y)" by (import hollight DEF_NUMLEFT) -constdefs - NUMRIGHT :: "nat => nat" +definition NUMRIGHT :: "nat => nat" where "NUMRIGHT == SOME Y::nat => nat. ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y" @@ -1691,8 +1668,7 @@ ALL (x::bool) y::nat. NUMLEFT (NUMSUM x y) = x & Y (NUMSUM x y) = y)" by (import hollight DEF_NUMRIGHT) -constdefs - INJN :: "nat => nat => 'A => bool" +definition INJN :: "nat => nat => 'A => bool" where "INJN == %(u::nat) (n::nat) a::'A::type. n = u" lemma DEF_INJN: "INJN = (%(u::nat) (n::nat) a::'A::type. n = u)" @@ -1710,8 +1686,7 @@ ((op =::nat => nat => bool) n1 n2)))" by (import hollight INJN_INJ) -constdefs - INJA :: "'A => nat => 'A => bool" +definition INJA :: "'A => nat => 'A => bool" where "INJA == %(u::'A::type) (n::nat) b::'A::type. b = u" lemma DEF_INJA: "INJA = (%(u::'A::type) (n::nat) b::'A::type. b = u)" @@ -1720,8 +1695,7 @@ lemma INJA_INJ: "ALL (a1::'A::type) a2::'A::type. (INJA a1 = INJA a2) = (a1 = a2)" by (import hollight INJA_INJ) -constdefs - INJF :: "(nat => nat => 'A => bool) => nat => 'A => bool" +definition INJF :: "(nat => nat => 'A => bool) => nat => 'A => bool" where "INJF == %(u::nat => nat => 'A::type => bool) n::nat. u (NUMFST n) (NUMSND n)" lemma DEF_INJF: "INJF = @@ -1732,8 +1706,7 @@ (INJF f1 = INJF f2) = (f1 = f2)" by (import hollight INJF_INJ) -constdefs - INJP :: "(nat => 'A => bool) => (nat => 'A => bool) => nat => 'A => bool" +definition INJP :: "(nat => 'A => bool) => (nat => 'A => bool) => nat => 'A => bool" where "INJP == %(u::nat => 'A::type => bool) (ua::nat => 'A::type => bool) (n::nat) a::'A::type. COND (NUMLEFT n) (u (NUMRIGHT n) a) (ua (NUMRIGHT n) a)" @@ -1748,8 +1721,7 @@ (INJP f1 f2 = INJP f1' f2') = (f1 = f1' & f2 = f2')" by (import hollight INJP_INJ) -constdefs - ZCONSTR :: "nat => 'A => (nat => nat => 'A => bool) => nat => 'A => bool" +definition ZCONSTR :: "nat => 'A => (nat => nat => 'A => bool) => nat => 'A => bool" where "ZCONSTR == %(u::nat) (ua::'A::type) ub::nat => nat => 'A::type => bool. INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub))" @@ -1759,8 +1731,7 @@ INJP (INJN (Suc u)) (INJP (INJA ua) (INJF ub)))" by (import hollight DEF_ZCONSTR) -constdefs - ZBOT :: "nat => 'A => bool" +definition ZBOT :: "nat => 'A => bool" where "ZBOT == INJP (INJN 0) (SOME z::nat => 'A::type => bool. True)" lemma DEF_ZBOT: "ZBOT = INJP (INJN 0) (SOME z::nat => 'A::type => bool. True)" @@ -1770,8 +1741,7 @@ ZCONSTR x xa xb ~= ZBOT" by (import hollight ZCONSTR_ZBOT) -constdefs - ZRECSPACE :: "(nat => 'A => bool) => bool" +definition ZRECSPACE :: "(nat => 'A => bool) => bool" where "ZRECSPACE == %a::nat => 'A::type => bool. ALL ZRECSPACE'::(nat => 'A::type => bool) => bool. @@ -1809,8 +1779,7 @@ [where a="a :: 'A recspace" and r=r , OF type_definition_recspace] -constdefs - BOTTOM :: "'A recspace" +definition BOTTOM :: "'A recspace" where "(op ==::'A::type recspace => 'A::type recspace => prop) (BOTTOM::'A::type recspace) ((_mk_rec::(nat => 'A::type => bool) => 'A::type recspace) @@ -1822,8 +1791,7 @@ (ZBOT::nat => 'A::type => bool))" by (import hollight DEF_BOTTOM) -constdefs - CONSTR :: "nat => 'A => (nat => 'A recspace) => 'A recspace" +definition CONSTR :: "nat => 'A => (nat => 'A recspace) => 'A recspace" where "(op ==::(nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) => (nat => 'A::type => (nat => 'A::type recspace) => 'A::type recspace) @@ -1900,8 +1868,7 @@ f (CONSTR c i r) = Fn c i r (%n::nat. f (r n))" by (import hollight CONSTR_REC) -constdefs - FCONS :: "'A => (nat => 'A) => nat => 'A" +definition FCONS :: "'A => (nat => 'A) => nat => 'A" where "FCONS == SOME FCONS::'A::type => (nat => 'A::type) => nat => 'A::type. (ALL (a::'A::type) f::nat => 'A::type. FCONS a f 0 = a) & @@ -1917,8 +1884,7 @@ lemma FCONS_UNDO: "ALL f::nat => 'A::type. f = FCONS (f 0) (f o Suc)" by (import hollight FCONS_UNDO) -constdefs - FNIL :: "nat => 'A" +definition FNIL :: "nat => 'A" where "FNIL == %u::nat. SOME x::'A::type. True" lemma DEF_FNIL: "FNIL = (%u::nat. SOME x::'A::type. True)" @@ -1995,8 +1961,7 @@ [where a="a :: 'A hollight.option" and r=r , OF type_definition_option] -constdefs - NONE :: "'A hollight.option" +definition NONE :: "'A hollight.option" where "(op ==::'A::type hollight.option => 'A::type hollight.option => prop) (NONE::'A::type hollight.option) ((_mk_option::'A::type recspace => 'A::type hollight.option) @@ -2093,8 +2058,7 @@ [where a="a :: 'A hollight.list" and r=r , OF type_definition_list] -constdefs - NIL :: "'A hollight.list" +definition NIL :: "'A hollight.list" where "(op ==::'A::type hollight.list => 'A::type hollight.list => prop) (NIL::'A::type hollight.list) ((_mk_list::'A::type recspace => 'A::type hollight.list) @@ -2114,8 +2078,7 @@ (%n::nat. BOTTOM::'A::type recspace)))" by (import hollight DEF_NIL) -constdefs - CONS :: "'A => 'A hollight.list => 'A hollight.list" +definition CONS :: "'A => 'A hollight.list => 'A hollight.list" where "(op ==::('A::type => 'A::type hollight.list => 'A::type hollight.list) => ('A::type => 'A::type hollight.list => 'A::type hollight.list) => prop) @@ -2160,8 +2123,7 @@ EX x::bool => 'A::type. x False = a & x True = b" by (import hollight bool_RECURSION) -constdefs - ISO :: "('A => 'B) => ('B => 'A) => bool" +definition ISO :: "('A => 'B) => ('B => 'A) => bool" where "ISO == %(u::'A::type => 'B::type) ua::'B::type => 'A::type. (ALL x::'B::type. u (ua x) = x) & (ALL y::'A::type. ua (u y) = y)" @@ -2244,15 +2206,13 @@ (%n::nat. BOTTOM::bool recspace)))" by (import hollight DEF__10303) -constdefs - two_1 :: "N_2" +definition two_1 :: "N_2" where "two_1 == _10302" lemma DEF_two_1: "two_1 = _10302" by (import hollight DEF_two_1) -constdefs - two_2 :: "N_2" +definition two_2 :: "N_2" where "two_2 == _10303" lemma DEF_two_2: "two_2 = _10303" @@ -2337,22 +2297,19 @@ (%n::nat. BOTTOM::bool recspace)))" by (import hollight DEF__10328) -constdefs - three_1 :: "N_3" +definition three_1 :: "N_3" where "three_1 == _10326" lemma DEF_three_1: "three_1 = _10326" by (import hollight DEF_three_1) -constdefs - three_2 :: "N_3" +definition three_2 :: "N_3" where "three_2 == _10327" lemma DEF_three_2: "three_2 = _10327" by (import hollight DEF_three_2) -constdefs - three_3 :: "N_3" +definition three_3 :: "N_3" where "three_3 == _10328" lemma DEF_three_3: "three_3 = _10328" @@ -2365,8 +2322,7 @@ All P" by (import hollight list_INDUCT) -constdefs - HD :: "'A hollight.list => 'A" +definition HD :: "'A hollight.list => 'A" where "HD == SOME HD::'A::type hollight.list => 'A::type. ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h" @@ -2376,8 +2332,7 @@ ALL (t::'A::type hollight.list) h::'A::type. HD (CONS h t) = h)" by (import hollight DEF_HD) -constdefs - TL :: "'A hollight.list => 'A hollight.list" +definition TL :: "'A hollight.list => 'A hollight.list" where "TL == SOME TL::'A::type hollight.list => 'A::type hollight.list. ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t" @@ -2387,8 +2342,7 @@ ALL (h::'A::type) t::'A::type hollight.list. TL (CONS h t) = t)" by (import hollight DEF_TL) -constdefs - APPEND :: "'A hollight.list => 'A hollight.list => 'A hollight.list" +definition APPEND :: "'A hollight.list => 'A hollight.list => 'A hollight.list" where "APPEND == SOME APPEND::'A::type hollight.list => 'A::type hollight.list => 'A::type hollight.list. @@ -2405,8 +2359,7 @@ APPEND (CONS h t) l = CONS h (APPEND t l)))" by (import hollight DEF_APPEND) -constdefs - REVERSE :: "'A hollight.list => 'A hollight.list" +definition REVERSE :: "'A hollight.list => 'A hollight.list" where "REVERSE == SOME REVERSE::'A::type hollight.list => 'A::type hollight.list. REVERSE NIL = NIL & @@ -2420,8 +2373,7 @@ REVERSE (CONS x l) = APPEND (REVERSE l) (CONS x NIL)))" by (import hollight DEF_REVERSE) -constdefs - LENGTH :: "'A hollight.list => nat" +definition LENGTH :: "'A hollight.list => nat" where "LENGTH == SOME LENGTH::'A::type hollight.list => nat. LENGTH NIL = 0 & @@ -2435,8 +2387,7 @@ LENGTH (CONS h t) = Suc (LENGTH t)))" by (import hollight DEF_LENGTH) -constdefs - MAP :: "('A => 'B) => 'A hollight.list => 'B hollight.list" +definition MAP :: "('A => 'B) => 'A hollight.list => 'B hollight.list" where "MAP == SOME MAP::('A::type => 'B::type) => 'A::type hollight.list => 'B::type hollight.list. @@ -2452,8 +2403,7 @@ MAP f (CONS h t) = CONS (f h) (MAP f t)))" by (import hollight DEF_MAP) -constdefs - LAST :: "'A hollight.list => 'A" +definition LAST :: "'A hollight.list => 'A" where "LAST == SOME LAST::'A::type hollight.list => 'A::type. ALL (h::'A::type) t::'A::type hollight.list. @@ -2465,8 +2415,7 @@ LAST (CONS h t) = COND (t = NIL) h (LAST t))" by (import hollight DEF_LAST) -constdefs - REPLICATE :: "nat => 'q_16860 => 'q_16860 hollight.list" +definition REPLICATE :: "nat => 'q_16860 => 'q_16860 hollight.list" where "REPLICATE == SOME REPLICATE::nat => 'q_16860::type => 'q_16860::type hollight.list. (ALL x::'q_16860::type. REPLICATE 0 x = NIL) & @@ -2480,8 +2429,7 @@ REPLICATE (Suc n) x = CONS x (REPLICATE n x)))" by (import hollight DEF_REPLICATE) -constdefs - NULL :: "'q_16875 hollight.list => bool" +definition NULL :: "'q_16875 hollight.list => bool" where "NULL == SOME NULL::'q_16875::type hollight.list => bool. NULL NIL = True & @@ -2495,8 +2443,7 @@ NULL (CONS h t) = False))" by (import hollight DEF_NULL) -constdefs - ALL_list :: "('q_16895 => bool) => 'q_16895 hollight.list => bool" +definition ALL_list :: "('q_16895 => bool) => 'q_16895 hollight.list => bool" where "ALL_list == SOME u::('q_16895::type => bool) => 'q_16895::type hollight.list => bool. (ALL P::'q_16895::type => bool. u P NIL = True) & @@ -2527,9 +2474,8 @@ t::'q_16916::type hollight.list. u P (CONS h t) = (P h | u P t)))" by (import hollight DEF_EX) -constdefs - ITLIST :: "('q_16939 => 'q_16938 => 'q_16938) -=> 'q_16939 hollight.list => 'q_16938 => 'q_16938" +definition ITLIST :: "('q_16939 => 'q_16938 => 'q_16938) +=> 'q_16939 hollight.list => 'q_16938 => 'q_16938" where "ITLIST == SOME ITLIST::('q_16939::type => 'q_16938::type => 'q_16938::type) => 'q_16939::type hollight.list @@ -2553,8 +2499,7 @@ ITLIST f (CONS h t) b = f h (ITLIST f t b)))" by (import hollight DEF_ITLIST) -constdefs - MEM :: "'q_16964 => 'q_16964 hollight.list => bool" +definition MEM :: "'q_16964 => 'q_16964 hollight.list => bool" where "MEM == SOME MEM::'q_16964::type => 'q_16964::type hollight.list => bool. (ALL x::'q_16964::type. MEM x NIL = False) & @@ -2570,9 +2515,8 @@ MEM x (CONS h t) = (x = h | MEM x t)))" by (import hollight DEF_MEM) -constdefs - ALL2 :: "('q_16997 => 'q_17004 => bool) -=> 'q_16997 hollight.list => 'q_17004 hollight.list => bool" +definition ALL2 :: "('q_16997 => 'q_17004 => bool) +=> 'q_16997 hollight.list => 'q_17004 hollight.list => bool" where "ALL2 == SOME ALL2::('q_16997::type => 'q_17004::type => bool) => 'q_16997::type hollight.list @@ -2604,10 +2548,9 @@ ALL2 P (CONS h1 t1) (CONS h2 t2) = (P h1 h2 & ALL2 P t1 t2)" by (import hollight ALL2) -constdefs - MAP2 :: "('q_17089 => 'q_17096 => 'q_17086) +definition MAP2 :: "('q_17089 => 'q_17096 => 'q_17086) => 'q_17089 hollight.list - => 'q_17096 hollight.list => 'q_17086 hollight.list" + => 'q_17096 hollight.list => 'q_17086 hollight.list" where "MAP2 == SOME MAP2::('q_17089::type => 'q_17096::type => 'q_17086::type) => 'q_17089::type hollight.list @@ -2639,8 +2582,7 @@ CONS (f h1 h2) (MAP2 f t1 t2)" by (import hollight MAP2) -constdefs - EL :: "nat => 'q_17157 hollight.list => 'q_17157" +definition EL :: "nat => 'q_17157 hollight.list => 'q_17157" where "EL == SOME EL::nat => 'q_17157::type hollight.list => 'q_17157::type. (ALL l::'q_17157::type hollight.list. EL 0 l = HD l) & @@ -2654,8 +2596,7 @@ EL (Suc n) l = EL n (TL l)))" by (import hollight DEF_EL) -constdefs - FILTER :: "('q_17182 => bool) => 'q_17182 hollight.list => 'q_17182 hollight.list" +definition FILTER :: "('q_17182 => bool) => 'q_17182 hollight.list => 'q_17182 hollight.list" where "FILTER == SOME FILTER::('q_17182::type => bool) => 'q_17182::type hollight.list @@ -2676,8 +2617,7 @@ COND (P h) (CONS h (FILTER P t)) (FILTER P t)))" by (import hollight DEF_FILTER) -constdefs - ASSOC :: "'q_17211 => ('q_17211 * 'q_17205) hollight.list => 'q_17205" +definition ASSOC :: "'q_17211 => ('q_17211 * 'q_17205) hollight.list => 'q_17205" where "ASSOC == SOME ASSOC::'q_17211::type => ('q_17211::type * 'q_17205::type) hollight.list @@ -2695,9 +2635,8 @@ ASSOC a (CONS h t) = COND (fst h = a) (snd h) (ASSOC a t))" by (import hollight DEF_ASSOC) -constdefs - ITLIST2 :: "('q_17235 => 'q_17243 => 'q_17233 => 'q_17233) -=> 'q_17235 hollight.list => 'q_17243 hollight.list => 'q_17233 => 'q_17233" +definition ITLIST2 :: "('q_17235 => 'q_17243 => 'q_17233 => 'q_17233) +=> 'q_17235 hollight.list => 'q_17243 hollight.list => 'q_17233 => 'q_17233" where "ITLIST2 == SOME ITLIST2::('q_17235::type => 'q_17243::type => 'q_17233::type => 'q_17233::type) @@ -3041,8 +2980,7 @@ ALL2 Q l l'" by (import hollight MONO_ALL2) -constdefs - dist :: "nat * nat => nat" +definition dist :: "nat * nat => nat" where "dist == %u::nat * nat. fst u - snd u + (snd u - fst u)" lemma DEF_dist: "dist = (%u::nat * nat. fst u - snd u + (snd u - fst u))" @@ -3133,8 +3071,7 @@ (EX (x::nat) N::nat. ALL i::nat. <= N i --> <= (P i) (Q i + x))" by (import hollight BOUNDS_IGNORE) -constdefs - is_nadd :: "(nat => nat) => bool" +definition is_nadd :: "(nat => nat) => bool" where "is_nadd == %u::nat => nat. EX B::nat. @@ -3211,8 +3148,7 @@ (A * n + B)" by (import hollight NADD_ALTMUL) -constdefs - nadd_eq :: "nadd => nadd => bool" +definition nadd_eq :: "nadd => nadd => bool" where "nadd_eq == %(u::nadd) ua::nadd. EX B::nat. ALL n::nat. <= (dist (dest_nadd u n, dest_nadd ua n)) B" @@ -3231,8 +3167,7 @@ lemma NADD_EQ_TRANS: "ALL (x::nadd) (y::nadd) z::nadd. nadd_eq x y & nadd_eq y z --> nadd_eq x z" by (import hollight NADD_EQ_TRANS) -constdefs - nadd_of_num :: "nat => nadd" +definition nadd_of_num :: "nat => nadd" where "nadd_of_num == %u::nat. mk_nadd (op * u)" lemma DEF_nadd_of_num: "nadd_of_num = (%u::nat. mk_nadd (op * u))" @@ -3247,8 +3182,7 @@ lemma NADD_OF_NUM_EQ: "ALL (m::nat) n::nat. nadd_eq (nadd_of_num m) (nadd_of_num n) = (m = n)" by (import hollight NADD_OF_NUM_EQ) -constdefs - nadd_le :: "nadd => nadd => bool" +definition nadd_le :: "nadd => nadd => bool" where "nadd_le == %(u::nadd) ua::nadd. EX B::nat. ALL n::nat. <= (dest_nadd u n) (dest_nadd ua n + B)" @@ -3289,8 +3223,7 @@ lemma NADD_OF_NUM_LE: "ALL (m::nat) n::nat. nadd_le (nadd_of_num m) (nadd_of_num n) = <= m n" by (import hollight NADD_OF_NUM_LE) -constdefs - nadd_add :: "nadd => nadd => nadd" +definition nadd_add :: "nadd => nadd => nadd" where "nadd_add == %(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u n + dest_nadd ua n)" @@ -3332,8 +3265,7 @@ (nadd_of_num (x + xa))" by (import hollight NADD_OF_NUM_ADD) -constdefs - nadd_mul :: "nadd => nadd => nadd" +definition nadd_mul :: "nadd => nadd => nadd" where "nadd_mul == %(u::nadd) ua::nadd. mk_nadd (%n::nat. dest_nadd u (dest_nadd ua n))" @@ -3438,8 +3370,7 @@ (EX (A::nat) N::nat. ALL n::nat. <= N n --> <= n (A * dest_nadd x n))" by (import hollight NADD_LBOUND) -constdefs - nadd_rinv :: "nadd => nat => nat" +definition nadd_rinv :: "nadd => nat => nat" where "nadd_rinv == %(u::nadd) n::nat. DIV (n * n) (dest_nadd u n)" lemma DEF_nadd_rinv: "nadd_rinv = (%(u::nadd) n::nat. DIV (n * n) (dest_nadd u n))" @@ -3528,8 +3459,7 @@ <= (dist (m * nadd_rinv x n, n * nadd_rinv x m)) (B * (m + n)))" by (import hollight NADD_MUL_LINV_LEMMA8) -constdefs - nadd_inv :: "nadd => nadd" +definition nadd_inv :: "nadd => nadd" where "nadd_inv == %u::nadd. COND (nadd_eq u (nadd_of_num 0)) (nadd_of_num 0) (mk_nadd (nadd_rinv u))" @@ -3570,15 +3500,13 @@ [where a="a :: hreal" and r=r , OF type_definition_hreal] -constdefs - hreal_of_num :: "nat => hreal" +definition hreal_of_num :: "nat => hreal" where "hreal_of_num == %m::nat. mk_hreal (nadd_eq (nadd_of_num m))" lemma DEF_hreal_of_num: "hreal_of_num = (%m::nat. mk_hreal (nadd_eq (nadd_of_num m)))" by (import hollight DEF_hreal_of_num) -constdefs - hreal_add :: "hreal => hreal => hreal" +definition hreal_add :: "hreal => hreal => hreal" where "hreal_add == %(x::hreal) y::hreal. mk_hreal @@ -3594,8 +3522,7 @@ nadd_eq (nadd_add xa ya) u & dest_hreal x xa & dest_hreal y ya))" by (import hollight DEF_hreal_add) -constdefs - hreal_mul :: "hreal => hreal => hreal" +definition hreal_mul :: "hreal => hreal => hreal" where "hreal_mul == %(x::hreal) y::hreal. mk_hreal @@ -3611,8 +3538,7 @@ nadd_eq (nadd_mul xa ya) u & dest_hreal x xa & dest_hreal y ya))" by (import hollight DEF_hreal_mul) -constdefs - hreal_le :: "hreal => hreal => bool" +definition hreal_le :: "hreal => hreal => bool" where "hreal_le == %(x::hreal) y::hreal. SOME u::bool. @@ -3626,8 +3552,7 @@ nadd_le xa ya = u & dest_hreal x xa & dest_hreal y ya)" by (import hollight DEF_hreal_le) -constdefs - hreal_inv :: "hreal => hreal" +definition hreal_inv :: "hreal => hreal" where "hreal_inv == %x::hreal. mk_hreal @@ -3685,22 +3610,19 @@ hreal_le a b --> hreal_le (hreal_mul a c) (hreal_mul b c)" by (import hollight HREAL_LE_MUL_RCANCEL_IMP) -constdefs - treal_of_num :: "nat => hreal * hreal" +definition treal_of_num :: "nat => hreal * hreal" where "treal_of_num == %u::nat. (hreal_of_num u, hreal_of_num 0)" lemma DEF_treal_of_num: "treal_of_num = (%u::nat. (hreal_of_num u, hreal_of_num 0))" by (import hollight DEF_treal_of_num) -constdefs - treal_neg :: "hreal * hreal => hreal * hreal" +definition treal_neg :: "hreal * hreal => hreal * hreal" where "treal_neg == %u::hreal * hreal. (snd u, fst u)" lemma DEF_treal_neg: "treal_neg = (%u::hreal * hreal. (snd u, fst u))" by (import hollight DEF_treal_neg) -constdefs - treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" +definition treal_add :: "hreal * hreal => hreal * hreal => hreal * hreal" where "treal_add == %(u::hreal * hreal) ua::hreal * hreal. (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua))" @@ -3710,8 +3632,7 @@ (hreal_add (fst u) (fst ua), hreal_add (snd u) (snd ua)))" by (import hollight DEF_treal_add) -constdefs - treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" +definition treal_mul :: "hreal * hreal => hreal * hreal => hreal * hreal" where "treal_mul == %(u::hreal * hreal) ua::hreal * hreal. (hreal_add (hreal_mul (fst u) (fst ua)) (hreal_mul (snd u) (snd ua)), @@ -3723,8 +3644,7 @@ hreal_add (hreal_mul (fst u) (snd ua)) (hreal_mul (snd u) (fst ua))))" by (import hollight DEF_treal_mul) -constdefs - treal_le :: "hreal * hreal => hreal * hreal => bool" +definition treal_le :: "hreal * hreal => hreal * hreal => bool" where "treal_le == %(u::hreal * hreal) ua::hreal * hreal. hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u))" @@ -3734,8 +3654,7 @@ hreal_le (hreal_add (fst u) (snd ua)) (hreal_add (fst ua) (snd u)))" by (import hollight DEF_treal_le) -constdefs - treal_inv :: "hreal * hreal => hreal * hreal" +definition treal_inv :: "hreal * hreal => hreal * hreal" where "treal_inv == %u::hreal * hreal. COND (fst u = snd u) (hreal_of_num 0, hreal_of_num 0) @@ -3755,8 +3674,7 @@ hreal_inv (SOME d::hreal. snd u = hreal_add (fst u) d))))" by (import hollight DEF_treal_inv) -constdefs - treal_eq :: "hreal * hreal => hreal * hreal => bool" +definition treal_eq :: "hreal * hreal => hreal * hreal => bool" where "treal_eq == %(u::hreal * hreal) ua::hreal * hreal. hreal_add (fst u) (snd ua) = hreal_add (fst ua) (snd u)" @@ -3916,15 +3834,13 @@ [where a="a :: hollight.real" and r=r , OF type_definition_real] -constdefs - real_of_num :: "nat => hollight.real" +definition real_of_num :: "nat => hollight.real" where "real_of_num == %m::nat. mk_real (treal_eq (treal_of_num m))" lemma DEF_real_of_num: "real_of_num = (%m::nat. mk_real (treal_eq (treal_of_num m)))" by (import hollight DEF_real_of_num) -constdefs - real_neg :: "hollight.real => hollight.real" +definition real_neg :: "hollight.real => hollight.real" where "real_neg == %x1::hollight.real. mk_real @@ -3940,8 +3856,7 @@ treal_eq (treal_neg x1a) u & dest_real x1 x1a))" by (import hollight DEF_real_neg) -constdefs - real_add :: "hollight.real => hollight.real => hollight.real" +definition real_add :: "hollight.real => hollight.real => hollight.real" where "real_add == %(x1::hollight.real) y1::hollight.real. mk_real @@ -3959,8 +3874,7 @@ dest_real x1 x1a & dest_real y1 y1a))" by (import hollight DEF_real_add) -constdefs - real_mul :: "hollight.real => hollight.real => hollight.real" +definition real_mul :: "hollight.real => hollight.real => hollight.real" where "real_mul == %(x1::hollight.real) y1::hollight.real. mk_real @@ -3978,8 +3892,7 @@ dest_real x1 x1a & dest_real y1 y1a))" by (import hollight DEF_real_mul) -constdefs - real_le :: "hollight.real => hollight.real => bool" +definition real_le :: "hollight.real => hollight.real => bool" where "real_le == %(x1::hollight.real) y1::hollight.real. SOME u::bool. @@ -3993,8 +3906,7 @@ treal_le x1a y1a = u & dest_real x1 x1a & dest_real y1 y1a)" by (import hollight DEF_real_le) -constdefs - real_inv :: "hollight.real => hollight.real" +definition real_inv :: "hollight.real => hollight.real" where "real_inv == %x::hollight.real. mk_real @@ -4008,15 +3920,13 @@ EX xa::hreal * hreal. treal_eq (treal_inv xa) u & dest_real x xa))" by (import hollight DEF_real_inv) -constdefs - real_sub :: "hollight.real => hollight.real => hollight.real" +definition real_sub :: "hollight.real => hollight.real => hollight.real" where "real_sub == %(u::hollight.real) ua::hollight.real. real_add u (real_neg ua)" lemma DEF_real_sub: "real_sub = (%(u::hollight.real) ua::hollight.real. real_add u (real_neg ua))" by (import hollight DEF_real_sub) -constdefs - real_lt :: "hollight.real => hollight.real => bool" +definition real_lt :: "hollight.real => hollight.real => bool" where "real_lt == %(u::hollight.real) ua::hollight.real. ~ real_le ua u" lemma DEF_real_lt: "real_lt = (%(u::hollight.real) ua::hollight.real. ~ real_le ua u)" @@ -4040,8 +3950,7 @@ lemma DEF_real_gt: "hollight.real_gt = (%(u::hollight.real) ua::hollight.real. real_lt ua u)" by (import hollight DEF_real_gt) -constdefs - real_abs :: "hollight.real => hollight.real" +definition real_abs :: "hollight.real => hollight.real" where "real_abs == %u::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u)" @@ -4049,8 +3958,7 @@ (%u::hollight.real. COND (real_le (real_of_num 0) u) u (real_neg u))" by (import hollight DEF_real_abs) -constdefs - real_pow :: "hollight.real => nat => hollight.real" +definition real_pow :: "hollight.real => nat => hollight.real" where "real_pow == SOME real_pow::hollight.real => nat => hollight.real. (ALL x::hollight.real. real_pow x 0 = real_of_num (NUMERAL_BIT1 0)) & @@ -4064,22 +3972,19 @@ real_pow x (Suc n) = real_mul x (real_pow x n)))" by (import hollight DEF_real_pow) -constdefs - real_div :: "hollight.real => hollight.real => hollight.real" +definition real_div :: "hollight.real => hollight.real => hollight.real" where "real_div == %(u::hollight.real) ua::hollight.real. real_mul u (real_inv ua)" lemma DEF_real_div: "real_div = (%(u::hollight.real) ua::hollight.real. real_mul u (real_inv ua))" by (import hollight DEF_real_div) -constdefs - real_max :: "hollight.real => hollight.real => hollight.real" +definition real_max :: "hollight.real => hollight.real => hollight.real" where "real_max == %(u::hollight.real) ua::hollight.real. COND (real_le u ua) ua u" lemma DEF_real_max: "real_max = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) ua u)" by (import hollight DEF_real_max) -constdefs - real_min :: "hollight.real => hollight.real => hollight.real" +definition real_min :: "hollight.real => hollight.real => hollight.real" where "real_min == %(u::hollight.real) ua::hollight.real. COND (real_le u ua) u ua" lemma DEF_real_min: "real_min = (%(u::hollight.real) ua::hollight.real. COND (real_le u ua) u ua)" @@ -5212,8 +5117,7 @@ (ALL x::hollight.real. All (P x))" by (import hollight REAL_WLOG_LT) -constdefs - mod_real :: "hollight.real => hollight.real => hollight.real => bool" +definition mod_real :: "hollight.real => hollight.real => hollight.real => bool" where "mod_real == %(u::hollight.real) (ua::hollight.real) ub::hollight.real. EX q::hollight.real. real_sub ua ub = real_mul q u" @@ -5223,8 +5127,7 @@ EX q::hollight.real. real_sub ua ub = real_mul q u)" by (import hollight DEF_mod_real) -constdefs - DECIMAL :: "nat => nat => hollight.real" +definition DECIMAL :: "nat => nat => hollight.real" where "DECIMAL == %(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua)" lemma DEF_DECIMAL: "DECIMAL = (%(u::nat) ua::nat. real_div (real_of_num u) (real_of_num ua))" @@ -5267,8 +5170,7 @@ (real_mul x1 y2 = real_mul x2 y1)" by (import hollight RAT_LEMMA5) -constdefs - is_int :: "hollight.real => bool" +definition is_int :: "hollight.real => bool" where "is_int == %u::hollight.real. EX n::nat. u = real_of_num n | u = real_neg (real_of_num n)" @@ -5297,8 +5199,7 @@ dest_int x = real_of_num n | dest_int x = real_neg (real_of_num n)" by (import hollight dest_int_rep) -constdefs - int_le :: "hollight.int => hollight.int => bool" +definition int_le :: "hollight.int => hollight.int => bool" where "int_le == %(u::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua)" @@ -5306,8 +5207,7 @@ (%(u::hollight.int) ua::hollight.int. real_le (dest_int u) (dest_int ua))" by (import hollight DEF_int_le) -constdefs - int_lt :: "hollight.int => hollight.int => bool" +definition int_lt :: "hollight.int => hollight.int => bool" where "int_lt == %(u::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua)" @@ -5315,8 +5215,7 @@ (%(u::hollight.int) ua::hollight.int. real_lt (dest_int u) (dest_int ua))" by (import hollight DEF_int_lt) -constdefs - int_ge :: "hollight.int => hollight.int => bool" +definition int_ge :: "hollight.int => hollight.int => bool" where "int_ge == %(u::hollight.int) ua::hollight.int. hollight.real_ge (dest_int u) (dest_int ua)" @@ -5326,8 +5225,7 @@ hollight.real_ge (dest_int u) (dest_int ua))" by (import hollight DEF_int_ge) -constdefs - int_gt :: "hollight.int => hollight.int => bool" +definition int_gt :: "hollight.int => hollight.int => bool" where "int_gt == %(u::hollight.int) ua::hollight.int. hollight.real_gt (dest_int u) (dest_int ua)" @@ -5337,8 +5235,7 @@ hollight.real_gt (dest_int u) (dest_int ua))" by (import hollight DEF_int_gt) -constdefs - int_of_num :: "nat => hollight.int" +definition int_of_num :: "nat => hollight.int" where "int_of_num == %u::nat. mk_int (real_of_num u)" lemma DEF_int_of_num: "int_of_num = (%u::nat. mk_int (real_of_num u))" @@ -5347,8 +5244,7 @@ lemma int_of_num_th: "ALL x::nat. dest_int (int_of_num x) = real_of_num x" by (import hollight int_of_num_th) -constdefs - int_neg :: "hollight.int => hollight.int" +definition int_neg :: "hollight.int => hollight.int" where "int_neg == %u::hollight.int. mk_int (real_neg (dest_int u))" lemma DEF_int_neg: "int_neg = (%u::hollight.int. mk_int (real_neg (dest_int u)))" @@ -5357,8 +5253,7 @@ lemma int_neg_th: "ALL x::hollight.int. dest_int (int_neg x) = real_neg (dest_int x)" by (import hollight int_neg_th) -constdefs - int_add :: "hollight.int => hollight.int => hollight.int" +definition int_add :: "hollight.int => hollight.int => hollight.int" where "int_add == %(u::hollight.int) ua::hollight.int. mk_int (real_add (dest_int u) (dest_int ua))" @@ -5372,8 +5267,7 @@ dest_int (int_add x xa) = real_add (dest_int x) (dest_int xa)" by (import hollight int_add_th) -constdefs - int_sub :: "hollight.int => hollight.int => hollight.int" +definition int_sub :: "hollight.int => hollight.int => hollight.int" where "int_sub == %(u::hollight.int) ua::hollight.int. mk_int (real_sub (dest_int u) (dest_int ua))" @@ -5387,8 +5281,7 @@ dest_int (int_sub x xa) = real_sub (dest_int x) (dest_int xa)" by (import hollight int_sub_th) -constdefs - int_mul :: "hollight.int => hollight.int => hollight.int" +definition int_mul :: "hollight.int => hollight.int => hollight.int" where "int_mul == %(u::hollight.int) ua::hollight.int. mk_int (real_mul (dest_int u) (dest_int ua))" @@ -5402,8 +5295,7 @@ dest_int (int_mul x y) = real_mul (dest_int x) (dest_int y)" by (import hollight int_mul_th) -constdefs - int_abs :: "hollight.int => hollight.int" +definition int_abs :: "hollight.int => hollight.int" where "int_abs == %u::hollight.int. mk_int (real_abs (dest_int u))" lemma DEF_int_abs: "int_abs = (%u::hollight.int. mk_int (real_abs (dest_int u)))" @@ -5412,8 +5304,7 @@ lemma int_abs_th: "ALL x::hollight.int. dest_int (int_abs x) = real_abs (dest_int x)" by (import hollight int_abs_th) -constdefs - int_max :: "hollight.int => hollight.int => hollight.int" +definition int_max :: "hollight.int => hollight.int => hollight.int" where "int_max == %(u::hollight.int) ua::hollight.int. mk_int (real_max (dest_int u) (dest_int ua))" @@ -5427,8 +5318,7 @@ dest_int (int_max x y) = real_max (dest_int x) (dest_int y)" by (import hollight int_max_th) -constdefs - int_min :: "hollight.int => hollight.int => hollight.int" +definition int_min :: "hollight.int => hollight.int => hollight.int" where "int_min == %(u::hollight.int) ua::hollight.int. mk_int (real_min (dest_int u) (dest_int ua))" @@ -5442,8 +5332,7 @@ dest_int (int_min x y) = real_min (dest_int x) (dest_int y)" by (import hollight int_min_th) -constdefs - int_pow :: "hollight.int => nat => hollight.int" +definition int_pow :: "hollight.int => nat => hollight.int" where "int_pow == %(u::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua)" lemma DEF_int_pow: "int_pow = (%(u::hollight.int) ua::nat. mk_int (real_pow (dest_int u) ua))" @@ -5496,8 +5385,7 @@ d ~= int_of_num 0 --> (EX c::hollight.int. int_lt x (int_mul c d))" by (import hollight INT_ARCH) -constdefs - mod_int :: "hollight.int => hollight.int => hollight.int => bool" +definition mod_int :: "hollight.int => hollight.int => hollight.int => bool" where "mod_int == %(u::hollight.int) (ua::hollight.int) ub::hollight.int. EX q::hollight.int. int_sub ua ub = int_mul q u" @@ -5507,8 +5395,7 @@ EX q::hollight.int. int_sub ua ub = int_mul q u)" by (import hollight DEF_mod_int) -constdefs - IN :: "'A => ('A => bool) => bool" +definition IN :: "'A => ('A => bool) => bool" where "IN == %(u::'A::type) ua::'A::type => bool. ua u" lemma DEF_IN: "IN = (%(u::'A::type) ua::'A::type => bool. ua u)" @@ -5518,15 +5405,13 @@ (x = xa) = (ALL xb::'A::type. IN xb x = IN xb xa)" by (import hollight EXTENSION) -constdefs - GSPEC :: "('A => bool) => 'A => bool" +definition GSPEC :: "('A => bool) => 'A => bool" where "GSPEC == %u::'A::type => bool. u" lemma DEF_GSPEC: "GSPEC = (%u::'A::type => bool. u)" by (import hollight DEF_GSPEC) -constdefs - SETSPEC :: "'q_37056 => bool => 'q_37056 => bool" +definition SETSPEC :: "'q_37056 => bool => 'q_37056 => bool" where "SETSPEC == %(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub" lemma DEF_SETSPEC: "SETSPEC = (%(u::'q_37056::type) (ua::bool) ub::'q_37056::type. ua & u = ub)" @@ -5548,15 +5433,13 @@ (ALL (p::'q_37194::type => bool) x::'q_37194::type. IN x p = p x)" by (import hollight IN_ELIM_THM) -constdefs - EMPTY :: "'A => bool" +definition EMPTY :: "'A => bool" where "EMPTY == %x::'A::type. False" lemma DEF_EMPTY: "EMPTY = (%x::'A::type. False)" by (import hollight DEF_EMPTY) -constdefs - INSERT :: "'A => ('A => bool) => 'A => bool" +definition INSERT :: "'A => ('A => bool) => 'A => bool" where "INSERT == %(u::'A::type) (ua::'A::type => bool) y::'A::type. IN y ua | y = u" lemma DEF_INSERT: "INSERT = @@ -5585,8 +5468,7 @@ GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u | IN x ua) x))" by (import hollight DEF_UNION) -constdefs - UNIONS :: "(('A => bool) => bool) => 'A => bool" +definition UNIONS :: "(('A => bool) => bool) => 'A => bool" where "UNIONS == %u::('A::type => bool) => bool. GSPEC @@ -5615,8 +5497,7 @@ GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & IN x ua) x))" by (import hollight DEF_INTER) -constdefs - INTERS :: "(('A => bool) => bool) => 'A => bool" +definition INTERS :: "(('A => bool) => bool) => 'A => bool" where "INTERS == %u::('A::type => bool) => bool. GSPEC @@ -5632,8 +5513,7 @@ SETSPEC ua (ALL ua::'A::type => bool. IN ua u --> IN x ua) x))" by (import hollight DEF_INTERS) -constdefs - DIFF :: "('A => bool) => ('A => bool) => 'A => bool" +definition DIFF :: "('A => bool) => ('A => bool) => 'A => bool" where "DIFF == %(u::'A::type => bool) ua::'A::type => bool. GSPEC (%ub::'A::type. EX x::'A::type. SETSPEC ub (IN x u & ~ IN x ua) x)" @@ -5648,8 +5528,7 @@ GSPEC (%u::'A::type. EX y::'A::type. SETSPEC u (IN y s | y = x) y)" by (import hollight INSERT) -constdefs - DELETE :: "('A => bool) => 'A => 'A => bool" +definition DELETE :: "('A => bool) => 'A => 'A => bool" where "DELETE == %(u::'A::type => bool) ua::'A::type. GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y)" @@ -5659,8 +5538,7 @@ GSPEC (%ub::'A::type. EX y::'A::type. SETSPEC ub (IN y u & y ~= ua) y))" by (import hollight DEF_DELETE) -constdefs - SUBSET :: "('A => bool) => ('A => bool) => bool" +definition SUBSET :: "('A => bool) => ('A => bool) => bool" where "SUBSET == %(u::'A::type => bool) ua::'A::type => bool. ALL x::'A::type. IN x u --> IN x ua" @@ -5670,8 +5548,7 @@ ALL x::'A::type. IN x u --> IN x ua)" by (import hollight DEF_SUBSET) -constdefs - PSUBSET :: "('A => bool) => ('A => bool) => bool" +definition PSUBSET :: "('A => bool) => ('A => bool) => bool" where "PSUBSET == %(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua" @@ -5679,8 +5556,7 @@ (%(u::'A::type => bool) ua::'A::type => bool. SUBSET u ua & u ~= ua)" by (import hollight DEF_PSUBSET) -constdefs - DISJOINT :: "('A => bool) => ('A => bool) => bool" +definition DISJOINT :: "('A => bool) => ('A => bool) => bool" where "DISJOINT == %(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY" @@ -5688,15 +5564,13 @@ (%(u::'A::type => bool) ua::'A::type => bool. hollight.INTER u ua = EMPTY)" by (import hollight DEF_DISJOINT) -constdefs - SING :: "('A => bool) => bool" +definition SING :: "('A => bool) => bool" where "SING == %u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY" lemma DEF_SING: "SING = (%u::'A::type => bool. EX x::'A::type. u = INSERT x EMPTY)" by (import hollight DEF_SING) -constdefs - FINITE :: "('A => bool) => bool" +definition FINITE :: "('A => bool) => bool" where "FINITE == %a::'A::type => bool. ALL FINITE'::('A::type => bool) => bool. @@ -5718,15 +5592,13 @@ FINITE' a)" by (import hollight DEF_FINITE) -constdefs - INFINITE :: "('A => bool) => bool" +definition INFINITE :: "('A => bool) => bool" where "INFINITE == %u::'A::type => bool. ~ FINITE u" lemma DEF_INFINITE: "INFINITE = (%u::'A::type => bool. ~ FINITE u)" by (import hollight DEF_INFINITE) -constdefs - IMAGE :: "('A => 'B) => ('A => bool) => 'B => bool" +definition IMAGE :: "('A => 'B) => ('A => bool) => 'B => bool" where "IMAGE == %(u::'A::type => 'B::type) ua::'A::type => bool. GSPEC @@ -5740,8 +5612,7 @@ EX y::'B::type. SETSPEC ub (EX x::'A::type. IN x ua & y = u x) y))" by (import hollight DEF_IMAGE) -constdefs - INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" +definition INJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where "INJ == %(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. (ALL x::'A::type. IN x ua --> IN (u x) ub) & @@ -5754,8 +5625,7 @@ IN x ua & IN y ua & u x = u y --> x = y))" by (import hollight DEF_INJ) -constdefs - SURJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" +definition SURJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where "SURJ == %(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. (ALL x::'A::type. IN x ua --> IN (u x) ub) & @@ -5767,8 +5637,7 @@ (ALL x::'B::type. IN x ub --> (EX y::'A::type. IN y ua & u y = x)))" by (import hollight DEF_SURJ) -constdefs - BIJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" +definition BIJ :: "('A => 'B) => ('A => bool) => ('B => bool) => bool" where "BIJ == %(u::'A::type => 'B::type) (ua::'A::type => bool) ub::'B::type => bool. INJ u ua ub & SURJ u ua ub" @@ -5778,22 +5647,19 @@ INJ u ua ub & SURJ u ua ub)" by (import hollight DEF_BIJ) -constdefs - CHOICE :: "('A => bool) => 'A" +definition CHOICE :: "('A => bool) => 'A" where "CHOICE == %u::'A::type => bool. SOME x::'A::type. IN x u" lemma DEF_CHOICE: "CHOICE = (%u::'A::type => bool. SOME x::'A::type. IN x u)" by (import hollight DEF_CHOICE) -constdefs - REST :: "('A => bool) => 'A => bool" +definition REST :: "('A => bool) => 'A => bool" where "REST == %u::'A::type => bool. DELETE u (CHOICE u)" lemma DEF_REST: "REST = (%u::'A::type => bool. DELETE u (CHOICE u))" by (import hollight DEF_REST) -constdefs - CARD_GE :: "('q_37693 => bool) => ('q_37690 => bool) => bool" +definition CARD_GE :: "('q_37693 => bool) => ('q_37690 => bool) => bool" where "CARD_GE == %(u::'q_37693::type => bool) ua::'q_37690::type => bool. EX f::'q_37693::type => 'q_37690::type. @@ -5807,8 +5673,7 @@ IN y ua --> (EX x::'q_37693::type. IN x u & y = f x))" by (import hollight DEF_CARD_GE) -constdefs - CARD_LE :: "('q_37702 => bool) => ('q_37701 => bool) => bool" +definition CARD_LE :: "('q_37702 => bool) => ('q_37701 => bool) => bool" where "CARD_LE == %(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u" @@ -5816,8 +5681,7 @@ (%(u::'q_37702::type => bool) ua::'q_37701::type => bool. CARD_GE ua u)" by (import hollight DEF_CARD_LE) -constdefs - CARD_EQ :: "('q_37712 => bool) => ('q_37713 => bool) => bool" +definition CARD_EQ :: "('q_37712 => bool) => ('q_37713 => bool) => bool" where "CARD_EQ == %(u::'q_37712::type => bool) ua::'q_37713::type => bool. CARD_LE u ua & CARD_LE ua u" @@ -5827,8 +5691,7 @@ CARD_LE u ua & CARD_LE ua u)" by (import hollight DEF_CARD_EQ) -constdefs - CARD_GT :: "('q_37727 => bool) => ('q_37728 => bool) => bool" +definition CARD_GT :: "('q_37727 => bool) => ('q_37728 => bool) => bool" where "CARD_GT == %(u::'q_37727::type => bool) ua::'q_37728::type => bool. CARD_GE u ua & ~ CARD_GE ua u" @@ -5838,8 +5701,7 @@ CARD_GE u ua & ~ CARD_GE ua u)" by (import hollight DEF_CARD_GT) -constdefs - CARD_LT :: "('q_37743 => bool) => ('q_37744 => bool) => bool" +definition CARD_LT :: "('q_37743 => bool) => ('q_37744 => bool) => bool" where "CARD_LT == %(u::'q_37743::type => bool) ua::'q_37744::type => bool. CARD_LE u ua & ~ CARD_LE ua u" @@ -5849,8 +5711,7 @@ CARD_LE u ua & ~ CARD_LE ua u)" by (import hollight DEF_CARD_LT) -constdefs - COUNTABLE :: "('q_37757 => bool) => bool" +definition COUNTABLE :: "('q_37757 => bool) => bool" where "(op ==::(('q_37757::type => bool) => bool) => (('q_37757::type => bool) => bool) => prop) (COUNTABLE::('q_37757::type => bool) => bool) @@ -6470,9 +6331,8 @@ FINITE s --> FINITE (DIFF s t)" by (import hollight FINITE_DIFF) -constdefs - FINREC :: "('q_41824 => 'q_41823 => 'q_41823) -=> 'q_41823 => ('q_41824 => bool) => 'q_41823 => nat => bool" +definition FINREC :: "('q_41824 => 'q_41823 => 'q_41823) +=> 'q_41823 => ('q_41824 => bool) => 'q_41823 => nat => bool" where "FINREC == SOME FINREC::('q_41824::type => 'q_41823::type => 'q_41823::type) => 'q_41823::type @@ -6558,9 +6418,8 @@ FINITE s --> g (INSERT x s) = COND (IN x s) (g s) (f x (g s))))" by (import hollight SET_RECURSION_LEMMA) -constdefs - ITSET :: "('q_42525 => 'q_42524 => 'q_42524) -=> ('q_42525 => bool) => 'q_42524 => 'q_42524" +definition ITSET :: "('q_42525 => 'q_42524 => 'q_42524) +=> ('q_42525 => bool) => 'q_42524 => 'q_42524" where "ITSET == %(u::'q_42525::type => 'q_42524::type => 'q_42524::type) (ua::'q_42525::type => bool) ub::'q_42524::type. @@ -6630,8 +6489,7 @@ EX x::'A::type. SETSPEC u (IN x s & (P::'A::type => bool) x) x))" by (import hollight FINITE_RESTRICT) -constdefs - CARD :: "('q_42918 => bool) => nat" +definition CARD :: "('q_42918 => bool) => nat" where "CARD == %u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0" lemma DEF_CARD: "CARD = (%u::'q_42918::type => bool. ITSET (%x::'q_42918::type. Suc) u 0)" @@ -6674,8 +6532,7 @@ CARD s + CARD t = CARD u" by (import hollight CARD_UNION_EQ) -constdefs - HAS_SIZE :: "('q_43199 => bool) => nat => bool" +definition HAS_SIZE :: "('q_43199 => bool) => nat => bool" where "HAS_SIZE == %(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua" lemma DEF_HAS_SIZE: "HAS_SIZE = (%(u::'q_43199::type => bool) ua::nat. FINITE u & CARD u = ua)" @@ -6944,8 +6801,7 @@ (ALL xa::'A::type. IN xa x --> (EX! m::nat. < m n & f m = xa)))" by (import hollight HAS_SIZE_INDEX) -constdefs - set_of_list :: "'q_45968 hollight.list => 'q_45968 => bool" +definition set_of_list :: "'q_45968 hollight.list => 'q_45968 => bool" where "set_of_list == SOME set_of_list::'q_45968::type hollight.list => 'q_45968::type => bool. set_of_list NIL = EMPTY & @@ -6959,8 +6815,7 @@ set_of_list (CONS h t) = INSERT h (set_of_list t)))" by (import hollight DEF_set_of_list) -constdefs - list_of_set :: "('q_45986 => bool) => 'q_45986 hollight.list" +definition list_of_set :: "('q_45986 => bool) => 'q_45986 hollight.list" where "list_of_set == %u::'q_45986::type => bool. SOME l::'q_45986::type hollight.list. @@ -6999,8 +6854,7 @@ hollight.UNION (set_of_list x) (set_of_list xa)" by (import hollight SET_OF_LIST_APPEND) -constdefs - pairwise :: "('q_46198 => 'q_46198 => bool) => ('q_46198 => bool) => bool" +definition pairwise :: "('q_46198 => 'q_46198 => bool) => ('q_46198 => bool) => bool" where "pairwise == %(u::'q_46198::type => 'q_46198::type => bool) ua::'q_46198::type => bool. ALL (x::'q_46198::type) y::'q_46198::type. @@ -7012,8 +6866,7 @@ IN x ua & IN y ua & x ~= y --> u x y)" by (import hollight DEF_pairwise) -constdefs - PAIRWISE :: "('q_46220 => 'q_46220 => bool) => 'q_46220 hollight.list => bool" +definition PAIRWISE :: "('q_46220 => 'q_46220 => bool) => 'q_46220 hollight.list => bool" where "PAIRWISE == SOME PAIRWISE::('q_46220::type => 'q_46220::type => bool) => 'q_46220::type hollight.list => bool. @@ -7075,8 +6928,7 @@ (EMPTY::'A::type => bool))))" by (import hollight FINITE_IMAGE_IMAGE) -constdefs - dimindex :: "('A => bool) => nat" +definition dimindex :: "('A => bool) => nat" where "(op ==::(('A::type => bool) => nat) => (('A::type => bool) => nat) => prop) (dimindex::('A::type => bool) => nat) (%u::'A::type => bool. @@ -7125,8 +6977,7 @@ dimindex s = dimindex t" by (import hollight DIMINDEX_FINITE_IMAGE) -constdefs - finite_index :: "nat => 'A" +definition finite_index :: "nat => 'A" where "(op ==::(nat => 'A::type) => (nat => 'A::type) => prop) (finite_index::nat => 'A::type) ((Eps::((nat => 'A::type) => bool) => nat => 'A::type) @@ -7287,8 +7138,7 @@ xa))))))" by (import hollight CART_EQ) -constdefs - lambda :: "(nat => 'A) => ('A, 'B) cart" +definition lambda :: "(nat => 'A) => ('A, 'B) cart" where "(op ==::((nat => 'A::type) => ('A::type, 'B::type) cart) => ((nat => 'A::type) => ('A::type, 'B::type) cart) => prop) (lambda::(nat => 'A::type) => ('A::type, 'B::type) cart) @@ -7388,8 +7238,7 @@ [where a="a :: ('A, 'B) finite_sum" and r=r , OF type_definition_finite_sum] -constdefs - pastecart :: "('A, 'M) cart => ('A, 'N) cart => ('A, ('M, 'N) finite_sum) cart" +definition pastecart :: "('A, 'M) cart => ('A, 'N) cart => ('A, ('M, 'N) finite_sum) cart" where "(op ==::(('A::type, 'M::type) cart => ('A::type, 'N::type) cart => ('A::type, ('M::type, 'N::type) finite_sum) cart) @@ -7439,8 +7288,7 @@ (hollight.UNIV::'M::type => bool))))))" by (import hollight DEF_pastecart) -constdefs - fstcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'M) cart" +definition fstcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'M) cart" where "fstcart == %u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u)" @@ -7448,8 +7296,7 @@ (%u::('A::type, ('M::type, 'N::type) finite_sum) cart. lambda ($ u))" by (import hollight DEF_fstcart) -constdefs - sndcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'N) cart" +definition sndcart :: "('A, ('M, 'N) finite_sum) cart => ('A, 'N) cart" where "(op ==::(('A::type, ('M::type, 'N::type) finite_sum) cart => ('A::type, 'N::type) cart) => (('A::type, ('M::type, 'N::type) finite_sum) cart @@ -7616,8 +7463,7 @@ (EX xb::'q_48070::type => 'q_48091::type. x = xb o xa)" by (import hollight FUNCTION_FACTORS_LEFT) -constdefs - dotdot :: "nat => nat => nat => bool" +definition dotdot :: "nat => nat => nat => bool" where "dotdot == %(u::nat) ua::nat. GSPEC (%ub::nat. EX x::nat. SETSPEC ub (<= u x & <= x ua) x)" @@ -7718,8 +7564,7 @@ SUBSET (dotdot m n) (dotdot p q) = (< n m | <= p m & <= n q)" by (import hollight SUBSET_NUMSEG) -constdefs - neutral :: "('q_48985 => 'q_48985 => 'q_48985) => 'q_48985" +definition neutral :: "('q_48985 => 'q_48985 => 'q_48985) => 'q_48985" where "neutral == %u::'q_48985::type => 'q_48985::type => 'q_48985::type. SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y" @@ -7729,8 +7574,7 @@ SOME x::'q_48985::type. ALL y::'q_48985::type. u x y = y & u y x = y)" by (import hollight DEF_neutral) -constdefs - monoidal :: "('A => 'A => 'A) => bool" +definition monoidal :: "('A => 'A => 'A) => bool" where "monoidal == %u::'A::type => 'A::type => 'A::type. (ALL (x::'A::type) y::'A::type. u x y = u y x) & @@ -7746,8 +7590,7 @@ (ALL x::'A::type. u (neutral u) x = x))" by (import hollight DEF_monoidal) -constdefs - support :: "('B => 'B => 'B) => ('A => 'B) => ('A => bool) => 'A => bool" +definition support :: "('B => 'B => 'B) => ('A => 'B) => ('A => bool) => 'A => bool" where "support == %(u::'B::type => 'B::type => 'B::type) (ua::'A::type => 'B::type) ub::'A::type => bool. @@ -7763,9 +7606,8 @@ EX x::'A::type. SETSPEC uc (IN x ub & ua x ~= neutral u) x))" by (import hollight DEF_support) -constdefs - iterate :: "('q_49090 => 'q_49090 => 'q_49090) -=> ('A => bool) => ('A => 'q_49090) => 'q_49090" +definition iterate :: "('q_49090 => 'q_49090 => 'q_49090) +=> ('A => bool) => ('A => 'q_49090) => 'q_49090" where "iterate == %(u::'q_49090::type => 'q_49090::type => 'q_49090::type) (ua::'A::type => bool) ub::'A::type => 'q_49090::type. @@ -8017,8 +7859,7 @@ iterate u_4247 s f = iterate u_4247 t g)" by (import hollight ITERATE_EQ_GENERAL) -constdefs - nsum :: "('q_51017 => bool) => ('q_51017 => nat) => nat" +definition nsum :: "('q_51017 => bool) => ('q_51017 => nat) => nat" where "(op ==::(('q_51017::type => bool) => ('q_51017::type => nat) => nat) => (('q_51017::type => bool) => ('q_51017::type => nat) => nat) => prop) @@ -8965,9 +8806,8 @@ hollight.sum x xb = hollight.sum xa xc" by (import hollight SUM_EQ_GENERAL) -constdefs - CASEWISE :: "(('q_57926 => 'q_57930) * ('q_57931 => 'q_57926 => 'q_57890)) hollight.list -=> 'q_57931 => 'q_57930 => 'q_57890" +definition CASEWISE :: "(('q_57926 => 'q_57930) * ('q_57931 => 'q_57926 => 'q_57890)) hollight.list +=> 'q_57931 => 'q_57930 => 'q_57890" where "CASEWISE == SOME CASEWISE::(('q_57926::type => 'q_57930::type) * ('q_57931::type @@ -9084,11 +8924,10 @@ x" by (import hollight CASEWISE_WORKS) -constdefs - admissible :: "('q_58228 => 'q_58221 => bool) +definition admissible :: "('q_58228 => 'q_58221 => bool) => (('q_58228 => 'q_58224) => 'q_58234 => bool) => ('q_58234 => 'q_58221) - => (('q_58228 => 'q_58224) => 'q_58234 => 'q_58229) => bool" + => (('q_58228 => 'q_58224) => 'q_58234 => 'q_58229) => bool" where "admissible == %(u::'q_58228::type => 'q_58221::type => bool) (ua::('q_58228::type => 'q_58224::type) => 'q_58234::type => bool) @@ -9114,10 +8953,9 @@ uc f a = uc g a)" by (import hollight DEF_admissible) -constdefs - tailadmissible :: "('A => 'A => bool) +definition tailadmissible :: "('A => 'A => bool) => (('A => 'B) => 'P => bool) - => ('P => 'A) => (('A => 'B) => 'P => 'B) => bool" + => ('P => 'A) => (('A => 'B) => 'P => 'B) => bool" where "tailadmissible == %(u::'A::type => 'A::type => bool) (ua::('A::type => 'B::type) => 'P::type => bool) @@ -9151,11 +8989,10 @@ ua f a --> uc f a = COND (P f a) (f (G f a)) (H f a)))" by (import hollight DEF_tailadmissible) -constdefs - superadmissible :: "('q_58378 => 'q_58378 => bool) +definition superadmissible :: "('q_58378 => 'q_58378 => bool) => (('q_58378 => 'q_58380) => 'q_58386 => bool) => ('q_58386 => 'q_58378) - => (('q_58378 => 'q_58380) => 'q_58386 => 'q_58380) => bool" + => (('q_58378 => 'q_58380) => 'q_58386 => 'q_58380) => bool" where "superadmissible == %(u::'q_58378::type => 'q_58378::type => bool) (ua::('q_58378::type => 'q_58380::type) => 'q_58386::type => bool) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Import/HOLLightCompat.thy --- a/src/HOL/Import/HOLLightCompat.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Import/HOLLightCompat.thy Mon Mar 01 13:42:31 2010 +0100 @@ -30,8 +30,7 @@ by simp qed -constdefs - Pred :: "nat \ nat" +definition Pred :: "nat \ nat" where "Pred n \ n - (Suc 0)" lemma Pred_altdef: "Pred = (SOME PRE. PRE 0 = 0 & (ALL n. PRE (Suc n) = n))" @@ -84,8 +83,7 @@ apply auto done -constdefs - NUMERAL_BIT0 :: "nat \ nat" +definition NUMERAL_BIT0 :: "nat \ nat" where "NUMERAL_BIT0 n \ n + n" lemma NUMERAL_BIT1_altdef: "NUMERAL_BIT1 n = Suc (n + n)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Isar_Examples/Expr_Compiler.thy --- a/src/HOL/Isar_Examples/Expr_Compiler.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Isar_Examples/Expr_Compiler.thy Mon Mar 01 13:42:31 2010 +0100 @@ -85,8 +85,7 @@ | Apply f => exec instrs (f (hd stack) (hd (tl stack)) # (tl (tl stack))) env)" -constdefs - execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val" +definition execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val" where "execute instrs env == hd (exec instrs [] env)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Isar_Examples/Hoare.thy --- a/src/HOL/Isar_Examples/Hoare.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Isar_Examples/Hoare.thy Mon Mar 01 13:42:31 2010 +0100 @@ -55,13 +55,10 @@ (if s : b then Sem c1 s s' else Sem c2 s s')" "Sem (While b x c) s s' = (EX n. iter n b (Sem c) s s')" -constdefs - Valid :: "'a bexp => 'a com => 'a bexp => bool" - ("(3|- _/ (2_)/ _)" [100, 55, 100] 50) +definition Valid :: "'a bexp => 'a com => 'a bexp => bool" ("(3|- _/ (2_)/ _)" [100, 55, 100] 50) where "|- P c Q == ALL s s'. Sem c s s' --> s : P --> s' : Q" -notation (xsymbols) - Valid ("(3\ _/ (2_)/ _)" [100, 55, 100] 50) +notation (xsymbols) Valid ("(3\ _/ (2_)/ _)" [100, 55, 100] 50) lemma ValidI [intro?]: "(!!s s'. Sem c s s' ==> s : P ==> s' : Q) ==> |- P c Q" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Isar_Examples/Mutilated_Checkerboard.thy --- a/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy Mon Mar 01 13:42:31 2010 +0100 @@ -60,8 +60,7 @@ subsection {* Basic properties of ``below'' *} -constdefs - below :: "nat => nat set" +definition below :: "nat => nat set" where "below n == {i. i < n}" lemma below_less_iff [iff]: "(i: below k) = (i < k)" @@ -84,8 +83,7 @@ subsection {* Basic properties of ``evnodd'' *} -constdefs - evnodd :: "(nat * nat) set => nat => (nat * nat) set" +definition evnodd :: "(nat * nat) set => nat => (nat * nat) set" where "evnodd A b == A Int {(i, j). (i + j) mod 2 = b}" lemma evnodd_iff: @@ -247,8 +245,7 @@ subsection {* Main theorem *} -constdefs - mutilated_board :: "nat => nat => (nat * nat) set" +definition mutilated_board :: "nat => nat => (nat * nat) set" where "mutilated_board m n == below (2 * (m + 1)) <*> below (2 * (n + 1)) - {(0, 0)} - {(2 * m + 1, 2 * n + 1)}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Matrix/ComputeNumeral.thy --- a/src/HOL/Matrix/ComputeNumeral.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Matrix/ComputeNumeral.thy Mon Mar 01 13:42:31 2010 +0100 @@ -20,7 +20,7 @@ lemmas bitiszero = iszero1 iszero2 iszero3 iszero4 (* lezero for bit strings *) -constdefs +definition "lezero x == (x \ 0)" lemma lezero1: "lezero Int.Pls = True" unfolding Int.Pls_def lezero_def by auto lemma lezero2: "lezero Int.Min = True" unfolding Int.Min_def lezero_def by auto @@ -60,7 +60,7 @@ lemmas bitarith = bitnorm bitiszero bitneg bitlezero biteq bitless bitle bitsucc bitpred bituminus bitadd bitmul -constdefs +definition "nat_norm_number_of (x::nat) == x" lemma nat_norm_number_of: "nat_norm_number_of (number_of w) = (if lezero w then 0 else number_of w)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Matrix/Matrix.thy --- a/src/HOL/Matrix/Matrix.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Matrix/Matrix.thy Mon Mar 01 13:42:31 2010 +0100 @@ -24,10 +24,10 @@ apply (rule Abs_matrix_induct) by (simp add: Abs_matrix_inverse matrix_def) -constdefs - nrows :: "('a::zero) matrix \ nat" +definition nrows :: "('a::zero) matrix \ nat" where "nrows A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image fst) (nonzero_positions (Rep_matrix A))))" - ncols :: "('a::zero) matrix \ nat" + +definition ncols :: "('a::zero) matrix \ nat" where "ncols A == if nonzero_positions(Rep_matrix A) = {} then 0 else Suc(Max ((image snd) (nonzero_positions (Rep_matrix A))))" lemma nrows: @@ -50,10 +50,10 @@ thus "Rep_matrix A m n = 0" by (simp add: nonzero_positions_def image_Collect) qed -constdefs - transpose_infmatrix :: "'a infmatrix \ 'a infmatrix" +definition transpose_infmatrix :: "'a infmatrix \ 'a infmatrix" where "transpose_infmatrix A j i == A i j" - transpose_matrix :: "('a::zero) matrix \ 'a matrix" + +definition transpose_matrix :: "('a::zero) matrix \ 'a matrix" where "transpose_matrix == Abs_matrix o transpose_infmatrix o Rep_matrix" declare transpose_infmatrix_def[simp] @@ -256,14 +256,16 @@ ultimately show "finite ?u" by (rule finite_subset) qed -constdefs - apply_infmatrix :: "('a \ 'b) \ 'a infmatrix \ 'b infmatrix" +definition apply_infmatrix :: "('a \ 'b) \ 'a infmatrix \ 'b infmatrix" where "apply_infmatrix f == % A. (% j i. f (A j i))" - apply_matrix :: "('a \ 'b) \ ('a::zero) matrix \ ('b::zero) matrix" + +definition apply_matrix :: "('a \ 'b) \ ('a::zero) matrix \ ('b::zero) matrix" where "apply_matrix f == % A. Abs_matrix (apply_infmatrix f (Rep_matrix A))" - combine_infmatrix :: "('a \ 'b \ 'c) \ 'a infmatrix \ 'b infmatrix \ 'c infmatrix" + +definition combine_infmatrix :: "('a \ 'b \ 'c) \ 'a infmatrix \ 'b infmatrix \ 'c infmatrix" where "combine_infmatrix f == % A B. (% j i. f (A j i) (B j i))" - combine_matrix :: "('a \ 'b \ 'c) \ ('a::zero) matrix \ ('b::zero) matrix \ ('c::zero) matrix" + +definition combine_matrix :: "('a \ 'b \ 'c) \ ('a::zero) matrix \ ('b::zero) matrix \ ('c::zero) matrix" where "combine_matrix f == % A B. Abs_matrix (combine_infmatrix f (Rep_matrix A) (Rep_matrix B))" lemma expand_apply_infmatrix[simp]: "apply_infmatrix f A j i = f (A j i)" @@ -272,10 +274,10 @@ lemma expand_combine_infmatrix[simp]: "combine_infmatrix f A B j i = f (A j i) (B j i)" by (simp add: combine_infmatrix_def) -constdefs -commutative :: "('a \ 'a \ 'b) \ bool" +definition commutative :: "('a \ 'a \ 'b) \ bool" where "commutative f == ! x y. f x y = f y x" -associative :: "('a \ 'a \ 'a) \ bool" + +definition associative :: "('a \ 'a \ 'a) \ bool" where "associative f == ! x y z. f (f x y) z = f x (f y z)" text{* @@ -356,12 +358,13 @@ lemma combine_ncols: "f 0 0 = 0 \ ncols A <= q \ ncols B <= q \ ncols(combine_matrix f A B) <= q" by (simp add: ncols_le) -constdefs - zero_r_neutral :: "('a \ 'b::zero \ 'a) \ bool" +definition zero_r_neutral :: "('a \ 'b::zero \ 'a) \ bool" where "zero_r_neutral f == ! a. f a 0 = a" - zero_l_neutral :: "('a::zero \ 'b \ 'b) \ bool" + +definition zero_l_neutral :: "('a::zero \ 'b \ 'b) \ bool" where "zero_l_neutral f == ! a. f 0 a = a" - zero_closed :: "(('a::zero) \ ('b::zero) \ ('c::zero)) \ bool" + +definition zero_closed :: "(('a::zero) \ ('b::zero) \ ('c::zero)) \ bool" where "zero_closed f == (!x. f x 0 = 0) & (!y. f 0 y = 0)" consts foldseq :: "('a \ 'a \ 'a) \ (nat \ 'a) \ nat \ 'a" @@ -648,10 +651,10 @@ then show ?concl by simp qed -constdefs - mult_matrix_n :: "nat \ (('a::zero) \ ('b::zero) \ ('c::zero)) \ ('c \ 'c \ 'c) \ 'a matrix \ 'b matrix \ 'c matrix" +definition mult_matrix_n :: "nat \ (('a::zero) \ ('b::zero) \ ('c::zero)) \ ('c \ 'c \ 'c) \ 'a matrix \ 'b matrix \ 'c matrix" where "mult_matrix_n n fmul fadd A B == Abs_matrix(% j i. foldseq fadd (% k. fmul (Rep_matrix A j k) (Rep_matrix B k i)) n)" - mult_matrix :: "(('a::zero) \ ('b::zero) \ ('c::zero)) \ ('c \ 'c \ 'c) \ 'a matrix \ 'b matrix \ 'c matrix" + +definition mult_matrix :: "(('a::zero) \ ('b::zero) \ ('c::zero)) \ ('c \ 'c \ 'c) \ 'a matrix \ 'b matrix \ 'c matrix" where "mult_matrix fmul fadd A B == mult_matrix_n (max (ncols A) (nrows B)) fmul fadd A B" lemma mult_matrix_n: @@ -673,12 +676,13 @@ finally show "mult_matrix_n n fmul fadd A B = mult_matrix_n m fmul fadd A B" by simp qed -constdefs - r_distributive :: "('a \ 'b \ 'b) \ ('b \ 'b \ 'b) \ bool" +definition r_distributive :: "('a \ 'b \ 'b) \ ('b \ 'b \ 'b) \ bool" where "r_distributive fmul fadd == ! a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)" - l_distributive :: "('a \ 'b \ 'a) \ ('a \ 'a \ 'a) \ bool" + +definition l_distributive :: "('a \ 'b \ 'a) \ ('a \ 'a \ 'a) \ bool" where "l_distributive fmul fadd == ! a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)" - distributive :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ bool" + +definition distributive :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ bool" where "distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd" lemma max1: "!! a x y. (a::nat) <= x \ a <= max x y" by (arith) @@ -835,20 +839,22 @@ apply (simp add: apply_matrix_def apply_infmatrix_def) by (simp add: zero_matrix_def) -constdefs - singleton_matrix :: "nat \ nat \ ('a::zero) \ 'a matrix" +definition singleton_matrix :: "nat \ nat \ ('a::zero) \ 'a matrix" where "singleton_matrix j i a == Abs_matrix(% m n. if j = m & i = n then a else 0)" - move_matrix :: "('a::zero) matrix \ int \ int \ 'a matrix" + +definition move_matrix :: "('a::zero) matrix \ int \ int \ 'a matrix" where "move_matrix A y x == Abs_matrix(% j i. if (neg ((int j)-y)) | (neg ((int i)-x)) then 0 else Rep_matrix A (nat ((int j)-y)) (nat ((int i)-x)))" - take_rows :: "('a::zero) matrix \ nat \ 'a matrix" + +definition take_rows :: "('a::zero) matrix \ nat \ 'a matrix" where "take_rows A r == Abs_matrix(% j i. if (j < r) then (Rep_matrix A j i) else 0)" - take_columns :: "('a::zero) matrix \ nat \ 'a matrix" + +definition take_columns :: "('a::zero) matrix \ nat \ 'a matrix" where "take_columns A c == Abs_matrix(% j i. if (i < c) then (Rep_matrix A j i) else 0)" -constdefs - column_of_matrix :: "('a::zero) matrix \ nat \ 'a matrix" +definition column_of_matrix :: "('a::zero) matrix \ nat \ 'a matrix" where "column_of_matrix A n == take_columns (move_matrix A 0 (- int n)) 1" - row_of_matrix :: "('a::zero) matrix \ nat \ 'a matrix" + +definition row_of_matrix :: "('a::zero) matrix \ nat \ 'a matrix" where "row_of_matrix A m == take_rows (move_matrix A (- int m) 0) 1" lemma Rep_singleton_matrix[simp]: "Rep_matrix (singleton_matrix j i e) m n = (if j = m & i = n then e else 0)" @@ -1042,10 +1048,10 @@ with contraprems show "False" by simp qed -constdefs - foldmatrix :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ ('a infmatrix) \ nat \ nat \ 'a" +definition foldmatrix :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ ('a infmatrix) \ nat \ nat \ 'a" where "foldmatrix f g A m n == foldseq_transposed g (% j. foldseq f (A j) n) m" - foldmatrix_transposed :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ ('a infmatrix) \ nat \ nat \ 'a" + +definition foldmatrix_transposed :: "('a \ 'a \ 'a) \ ('a \ 'a \ 'a) \ ('a infmatrix) \ nat \ nat \ 'a" where "foldmatrix_transposed f g A m n == foldseq g (% j. foldseq_transposed f (A j) n) m" lemma foldmatrix_transpose: @@ -1691,12 +1697,13 @@ lemma transpose_matrix_minus: "transpose_matrix (-(A::('a::group_add) matrix)) = - transpose_matrix (A::'a matrix)" by (simp add: minus_matrix_def transpose_apply_matrix) -constdefs - right_inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" +definition right_inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" where "right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X))) \ nrows X \ ncols A" - left_inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" + +definition left_inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" where "left_inverse_matrix A X == (X * A = one_matrix (max(nrows X) (ncols A))) \ ncols X \ nrows A" - inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" + +definition inverse_matrix :: "('a::{ring_1}) matrix \ 'a matrix \ bool" where "inverse_matrix A X == (right_inverse_matrix A X) \ (left_inverse_matrix A X)" lemma right_inverse_matrix_dim: "right_inverse_matrix A X \ nrows A = ncols X" @@ -1781,8 +1788,7 @@ lemma move_matrix_mult: "move_matrix ((A::('a::semiring_0) matrix)*B) j i = (move_matrix A j 0) * (move_matrix B 0 i)" by (simp add: move_matrix_ortho[of "A*B"] move_matrix_col_mult move_matrix_row_mult) -constdefs - scalar_mult :: "('a::ring) \ 'a matrix \ 'a matrix" +definition scalar_mult :: "('a::ring) \ 'a matrix \ 'a matrix" where "scalar_mult a m == apply_matrix (op * a) m" lemma scalar_mult_zero[simp]: "scalar_mult y 0 = 0" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Matrix/SparseMatrix.thy --- a/src/HOL/Matrix/SparseMatrix.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Matrix/SparseMatrix.thy Mon Mar 01 13:42:31 2010 +0100 @@ -552,8 +552,7 @@ else if j < i then (le_spvec [] b & le_spmat ((i,a)#as) bs) else (le_spvec a b & le_spmat as bs))" -constdefs - disj_matrices :: "('a::zero) matrix \ 'a matrix \ bool" +definition disj_matrices :: "('a::zero) matrix \ 'a matrix \ bool" where "disj_matrices A B == (! j i. (Rep_matrix A j i \ 0) \ (Rep_matrix B j i = 0)) & (! j i. (Rep_matrix B j i \ 0) \ (Rep_matrix A j i = 0))" declare [[simp_depth_limit = 6]] @@ -802,8 +801,7 @@ apply (simp_all add: sorted_spvec_abs_spvec) done -constdefs - diff_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat" +definition diff_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat" where "diff_spmat A B == add_spmat A (minus_spmat B)" lemma sorted_spmat_diff_spmat: "sorted_spmat A \ sorted_spmat B \ sorted_spmat (diff_spmat A B)" @@ -815,8 +813,7 @@ lemma sparse_row_diff_spmat: "sparse_row_matrix (diff_spmat A B ) = (sparse_row_matrix A) - (sparse_row_matrix B)" by (simp add: diff_spmat_def sparse_row_add_spmat sparse_row_matrix_minus) -constdefs - sorted_sparse_matrix :: "'a spmat \ bool" +definition sorted_sparse_matrix :: "'a spmat \ bool" where "sorted_sparse_matrix A == (sorted_spvec A) & (sorted_spmat A)" lemma sorted_sparse_matrix_imp_spvec: "sorted_sparse_matrix A \ sorted_spvec A" @@ -1014,8 +1011,7 @@ apply (simp_all add: sorted_nprt_spvec) done -constdefs - mult_est_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat \ 'a spmat \ 'a spmat" +definition mult_est_spmat :: "('a::lattice_ring) spmat \ 'a spmat \ 'a spmat \ 'a spmat \ 'a spmat" where "mult_est_spmat r1 r2 s1 s2 == add_spmat (mult_spmat (pprt_spmat s2) (pprt_spmat r2)) (add_spmat (mult_spmat (pprt_spmat s1) (nprt_spmat r2)) (add_spmat (mult_spmat (nprt_spmat s2) (pprt_spmat r1)) (mult_spmat (nprt_spmat s1) (nprt_spmat r1))))" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Metis_Examples/BigO.thy --- a/src/HOL/Metis_Examples/BigO.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Metis_Examples/BigO.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1099,9 +1099,7 @@ subsection {* Less than or equal to *} -constdefs - lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" - (infixl " 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl " ALL x. abs (g x) <= abs (f x) ==> diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Metis_Examples/Message.thy --- a/src/HOL/Metis_Examples/Message.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Metis_Examples/Message.thy Mon Mar 01 13:42:31 2010 +0100 @@ -26,8 +26,7 @@ text{*The inverse of a symmetric key is itself; that of a public key is the private key and vice versa*} -constdefs - symKeys :: "key set" +definition symKeys :: "key set" where "symKeys == {K. invKey K = K}" datatype --{*We allow any number of friendly agents*} @@ -55,12 +54,11 @@ "{|x, y|}" == "CONST MPair x y" -constdefs - HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) +definition HPair :: "[msg,msg] => msg" ("(4Hash[_] /_)" [0, 1000]) where --{*Message Y paired with a MAC computed with the help of X*} "Hash[X] Y == {| Hash{|X,Y|}, Y|}" - keysFor :: "msg set => key set" +definition keysFor :: "msg set => key set" where --{*Keys useful to decrypt elements of a message set*} "keysFor H == invKey ` {K. \X. Crypt K X \ H}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Metis_Examples/Tarski.thy --- a/src/HOL/Metis_Examples/Tarski.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Metis_Examples/Tarski.thy Mon Mar 01 13:42:31 2010 +0100 @@ -20,59 +20,56 @@ pset :: "'a set" order :: "('a * 'a) set" -constdefs - monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool" +definition monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool" where "monotone f A r == \x\A. \y\A. (x, y): r --> ((f x), (f y)) : r" - least :: "['a => bool, 'a potype] => 'a" +definition least :: "['a => bool, 'a potype] => 'a" where "least P po == @ x. x: pset po & P x & (\y \ pset po. P y --> (x,y): order po)" - greatest :: "['a => bool, 'a potype] => 'a" +definition greatest :: "['a => bool, 'a potype] => 'a" where "greatest P po == @ x. x: pset po & P x & (\y \ pset po. P y --> (y,x): order po)" - lub :: "['a set, 'a potype] => 'a" +definition lub :: "['a set, 'a potype] => 'a" where "lub S po == least (%x. \y\S. (y,x): order po) po" - glb :: "['a set, 'a potype] => 'a" +definition glb :: "['a set, 'a potype] => 'a" where "glb S po == greatest (%x. \y\S. (x,y): order po) po" - isLub :: "['a set, 'a potype, 'a] => bool" +definition isLub :: "['a set, 'a potype, 'a] => bool" where "isLub S po == %L. (L: pset po & (\y\S. (y,L): order po) & (\z\pset po. (\y\S. (y,z): order po) --> (L,z): order po))" - isGlb :: "['a set, 'a potype, 'a] => bool" +definition isGlb :: "['a set, 'a potype, 'a] => bool" where "isGlb S po == %G. (G: pset po & (\y\S. (G,y): order po) & (\z \ pset po. (\y\S. (z,y): order po) --> (z,G): order po))" - "fix" :: "[('a => 'a), 'a set] => 'a set" +definition "fix" :: "[('a => 'a), 'a set] => 'a set" where "fix f A == {x. x: A & f x = x}" - interval :: "[('a*'a) set,'a, 'a ] => 'a set" +definition interval :: "[('a*'a) set,'a, 'a ] => 'a set" where "interval r a b == {x. (a,x): r & (x,b): r}" -constdefs - Bot :: "'a potype => 'a" +definition Bot :: "'a potype => 'a" where "Bot po == least (%x. True) po" - Top :: "'a potype => 'a" +definition Top :: "'a potype => 'a" where "Top po == greatest (%x. True) po" - PartialOrder :: "('a potype) set" +definition PartialOrder :: "('a potype) set" where "PartialOrder == {P. refl_on (pset P) (order P) & antisym (order P) & trans (order P)}" - CompleteLattice :: "('a potype) set" +definition CompleteLattice :: "('a potype) set" where "CompleteLattice == {cl. cl: PartialOrder & (\S. S \ pset cl --> (\L. isLub S cl L)) & (\S. S \ pset cl --> (\G. isGlb S cl G))}" - induced :: "['a set, ('a * 'a) set] => ('a *'a)set" +definition induced :: "['a set, ('a * 'a) set] => ('a *'a)set" where "induced A r == {(a,b). a : A & b: A & (a,b): r}" -constdefs - sublattice :: "('a potype * 'a set)set" +definition sublattice :: "('a potype * 'a set)set" where "sublattice == SIGMA cl: CompleteLattice. {S. S \ pset cl & @@ -82,8 +79,7 @@ sublattice_syntax :: "['a set, 'a potype] => bool" ("_ <<= _" [51, 50] 50) where "S <<= cl \ S : sublattice `` {cl}" -constdefs - dual :: "'a potype => 'a potype" +definition dual :: "'a potype => 'a potype" where "dual po == (| pset = pset po, order = converse (order po) |)" locale PO = diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/Altern.thy --- a/src/HOL/MicroJava/BV/Altern.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/BV/Altern.thy Mon Mar 01 13:42:31 2010 +0100 @@ -8,19 +8,18 @@ imports BVSpec begin -constdefs - check_type :: "jvm_prog \ nat \ nat \ JVMType.state \ bool" +definition check_type :: "jvm_prog \ nat \ nat \ JVMType.state \ bool" where "check_type G mxs mxr s \ s \ states G mxs mxr" - wt_instr_altern :: "[instr,jvm_prog,ty,method_type,nat,nat,p_count, - exception_table,p_count] \ bool" +definition wt_instr_altern :: "[instr,jvm_prog,ty,method_type,nat,nat,p_count, + exception_table,p_count] \ bool" where "wt_instr_altern i G rT phi mxs mxr max_pc et pc \ app i G mxs rT pc et (phi!pc) \ check_type G mxs mxr (OK (phi!pc)) \ (\(pc',s') \ set (eff i G pc et (phi!pc)). pc' < max_pc \ G \ s' <=' phi!pc')" - wt_method_altern :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list, - exception_table,method_type] \ bool" +definition wt_method_altern :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list, + exception_table,method_type] \ bool" where "wt_method_altern G C pTs rT mxs mxl ins et phi \ let max_pc = length ins in 0 < max_pc \ diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/BVExample.thy --- a/src/HOL/MicroJava/BV/BVExample.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/BV/BVExample.thy Mon Mar 01 13:42:31 2010 +0100 @@ -167,8 +167,7 @@ @{prop [display] "P n"} *} -constdefs - intervall :: "nat \ nat \ nat \ bool" ("_ \ [_, _')") +definition intervall :: "nat \ nat \ nat \ bool" ("_ \ [_, _')") where "x \ [a, b) \ a \ x \ x < b" lemma pc_0: "x < n \ (x \ [0, n) \ P x) \ P x" @@ -238,8 +237,7 @@ lemmas eff_simps [simp] = eff_def norm_eff_def xcpt_eff_def declare appInvoke [simp del] -constdefs - phi_append :: method_type ("\\<^sub>a") +definition phi_append :: method_type ("\\<^sub>a") where "\\<^sub>a \ map (\(x,y). Some (x, map OK y)) [ ( [], [Class list_name, Class list_name]), ( [Class list_name], [Class list_name, Class list_name]), @@ -301,8 +299,7 @@ abbreviation Ctest :: ty where "Ctest == Class test_name" -constdefs - phi_makelist :: method_type ("\\<^sub>m") +definition phi_makelist :: method_type ("\\<^sub>m") where "\\<^sub>m \ map (\(x,y). Some (x, y)) [ ( [], [OK Ctest, Err , Err ]), ( [Clist], [OK Ctest, Err , Err ]), @@ -376,8 +373,7 @@ done text {* The whole program is welltyped: *} -constdefs - Phi :: prog_type ("\") +definition Phi :: prog_type ("\") where "\ C sg \ if C = test_name \ sg = (makelist_name, []) then \\<^sub>m else if C = list_name \ sg = (append_name, [Class list_name]) then \\<^sub>a else []" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/BVSpec.thy diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/Correct.thy --- a/src/HOL/MicroJava/BV/Correct.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/BV/Correct.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,47 +9,40 @@ imports BVSpec "../JVM/JVMExec" begin -constdefs - approx_val :: "[jvm_prog,aheap,val,ty err] \ bool" +definition approx_val :: "[jvm_prog,aheap,val,ty err] \ bool" where "approx_val G h v any == case any of Err \ True | OK T \ G,h\v::\T" - approx_loc :: "[jvm_prog,aheap,val list,locvars_type] \ bool" +definition approx_loc :: "[jvm_prog,aheap,val list,locvars_type] \ bool" where "approx_loc G hp loc LT == list_all2 (approx_val G hp) loc LT" - approx_stk :: "[jvm_prog,aheap,opstack,opstack_type] \ bool" +definition approx_stk :: "[jvm_prog,aheap,opstack,opstack_type] \ bool" where "approx_stk G hp stk ST == approx_loc G hp stk (map OK ST)" - correct_frame :: "[jvm_prog,aheap,state_type,nat,bytecode] \ frame \ bool" +definition correct_frame :: "[jvm_prog,aheap,state_type,nat,bytecode] \ frame \ bool" where "correct_frame G hp == \(ST,LT) maxl ins (stk,loc,C,sig,pc). approx_stk G hp stk ST \ approx_loc G hp loc LT \ pc < length ins \ length loc=length(snd sig)+maxl+1" - -consts - correct_frames :: "[jvm_prog,aheap,prog_type,ty,sig,frame list] \ bool" -primrec -"correct_frames G hp phi rT0 sig0 [] = True" +primrec correct_frames :: "[jvm_prog,aheap,prog_type,ty,sig,frame list] \ bool" where + "correct_frames G hp phi rT0 sig0 [] = True" +| "correct_frames G hp phi rT0 sig0 (f#frs) = + (let (stk,loc,C,sig,pc) = f in + (\ST LT rT maxs maxl ins et. + phi C sig ! pc = Some (ST,LT) \ is_class G C \ + method (G,C) sig = Some(C,rT,(maxs,maxl,ins,et)) \ + (\C' mn pTs. ins!pc = (Invoke C' mn pTs) \ + (mn,pTs) = sig0 \ + (\apTs D ST' LT'. + (phi C sig)!pc = Some ((rev apTs) @ (Class D) # ST', LT') \ + length apTs = length pTs \ + (\D' rT' maxs' maxl' ins' et'. + method (G,D) sig0 = Some(D',rT',(maxs',maxl',ins',et')) \ + G \ rT0 \ rT') \ + correct_frame G hp (ST, LT) maxl ins f \ + correct_frames G hp phi rT sig frs))))" -"correct_frames G hp phi rT0 sig0 (f#frs) = - (let (stk,loc,C,sig,pc) = f in - (\ST LT rT maxs maxl ins et. - phi C sig ! pc = Some (ST,LT) \ is_class G C \ - method (G,C) sig = Some(C,rT,(maxs,maxl,ins,et)) \ - (\C' mn pTs. ins!pc = (Invoke C' mn pTs) \ - (mn,pTs) = sig0 \ - (\apTs D ST' LT'. - (phi C sig)!pc = Some ((rev apTs) @ (Class D) # ST', LT') \ - length apTs = length pTs \ - (\D' rT' maxs' maxl' ins' et'. - method (G,D) sig0 = Some(D',rT',(maxs',maxl',ins',et')) \ - G \ rT0 \ rT') \ - correct_frame G hp (ST, LT) maxl ins f \ - correct_frames G hp phi rT sig frs))))" - - -constdefs - correct_state :: "[jvm_prog,prog_type,jvm_state] \ bool" - ("_,_ |-JVM _ [ok]" [51,51] 50) +definition correct_state :: "[jvm_prog,prog_type,jvm_state] \ bool" + ("_,_ |-JVM _ [ok]" [51,51] 50) where "correct_state G phi == \(xp,hp,frs). case xp of None \ (case frs of diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/Effect.thy --- a/src/HOL/MicroJava/BV/Effect.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/BV/Effect.thy Mon Mar 01 13:42:31 2010 +0100 @@ -13,27 +13,25 @@ succ_type = "(p_count \ state_type option) list" text {* Program counter of successor instructions: *} -consts - succs :: "instr \ p_count \ p_count list" -primrec +primrec succs :: "instr \ p_count \ p_count list" where "succs (Load idx) pc = [pc+1]" - "succs (Store idx) pc = [pc+1]" - "succs (LitPush v) pc = [pc+1]" - "succs (Getfield F C) pc = [pc+1]" - "succs (Putfield F C) pc = [pc+1]" - "succs (New C) pc = [pc+1]" - "succs (Checkcast C) pc = [pc+1]" - "succs Pop pc = [pc+1]" - "succs Dup pc = [pc+1]" - "succs Dup_x1 pc = [pc+1]" - "succs Dup_x2 pc = [pc+1]" - "succs Swap pc = [pc+1]" - "succs IAdd pc = [pc+1]" - "succs (Ifcmpeq b) pc = [pc+1, nat (int pc + b)]" - "succs (Goto b) pc = [nat (int pc + b)]" - "succs Return pc = [pc]" - "succs (Invoke C mn fpTs) pc = [pc+1]" - "succs Throw pc = [pc]" +| "succs (Store idx) pc = [pc+1]" +| "succs (LitPush v) pc = [pc+1]" +| "succs (Getfield F C) pc = [pc+1]" +| "succs (Putfield F C) pc = [pc+1]" +| "succs (New C) pc = [pc+1]" +| "succs (Checkcast C) pc = [pc+1]" +| "succs Pop pc = [pc+1]" +| "succs Dup pc = [pc+1]" +| "succs Dup_x1 pc = [pc+1]" +| "succs Dup_x2 pc = [pc+1]" +| "succs Swap pc = [pc+1]" +| "succs IAdd pc = [pc+1]" +| "succs (Ifcmpeq b) pc = [pc+1, nat (int pc + b)]" +| "succs (Goto b) pc = [nat (int pc + b)]" +| "succs Return pc = [pc]" +| "succs (Invoke C mn fpTs) pc = [pc+1]" +| "succs Throw pc = [pc]" text "Effect of instruction on the state type:" consts @@ -63,21 +61,16 @@ "eff' (Invoke C mn fpTs, G, (ST,LT)) = (let ST' = drop (length fpTs) ST in (fst (snd (the (method (G,C) (mn,fpTs))))#(tl ST'),LT))" - -consts - match_any :: "jvm_prog \ p_count \ exception_table \ cname list" -primrec +primrec match_any :: "jvm_prog \ p_count \ exception_table \ cname list" where "match_any G pc [] = []" - "match_any G pc (e#es) = (let (start_pc, end_pc, handler_pc, catch_type) = e; +| "match_any G pc (e#es) = (let (start_pc, end_pc, handler_pc, catch_type) = e; es' = match_any G pc es in if start_pc <= pc \ pc < end_pc then catch_type#es' else es')" -consts - match :: "jvm_prog \ xcpt \ p_count \ exception_table \ cname list" -primrec +primrec match :: "jvm_prog \ xcpt \ p_count \ exception_table \ cname list" where "match G X pc [] = []" - "match G X pc (e#es) = +| "match G X pc (e#es) = (if match_exception_entry G (Xcpt X) pc e then [Xcpt X] else match G X pc es)" lemma match_some_entry: @@ -96,23 +89,21 @@ "xcpt_names (i, G, pc, et) = []" -constdefs - xcpt_eff :: "instr \ jvm_prog \ p_count \ state_type option \ exception_table \ succ_type" +definition xcpt_eff :: "instr \ jvm_prog \ p_count \ state_type option \ exception_table \ succ_type" where "xcpt_eff i G pc s et == map (\C. (the (match_exception_table G C pc et), case s of None \ None | Some s' \ Some ([Class C], snd s'))) (xcpt_names (i,G,pc,et))" - norm_eff :: "instr \ jvm_prog \ state_type option \ state_type option" +definition norm_eff :: "instr \ jvm_prog \ state_type option \ state_type option" where "norm_eff i G == Option.map (\s. eff' (i,G,s))" - eff :: "instr \ jvm_prog \ p_count \ exception_table \ state_type option \ succ_type" +definition eff :: "instr \ jvm_prog \ p_count \ exception_table \ state_type option \ succ_type" where "eff i G pc et s == (map (\pc'. (pc',norm_eff i G s)) (succs i pc)) @ (xcpt_eff i G pc s et)" -constdefs - isPrimT :: "ty \ bool" +definition isPrimT :: "ty \ bool" where "isPrimT T == case T of PrimT T' \ True | RefT T' \ False" - isRefT :: "ty \ bool" +definition isRefT :: "ty \ bool" where "isRefT T == case T of PrimT T' \ False | RefT T' \ True" lemma isPrimT [simp]: @@ -177,11 +168,10 @@ "app' (i,G, pc,maxs,rT,s) = False" -constdefs - xcpt_app :: "instr \ jvm_prog \ nat \ exception_table \ bool" +definition xcpt_app :: "instr \ jvm_prog \ nat \ exception_table \ bool" where "xcpt_app i G pc et \ \C\set(xcpt_names (i,G,pc,et)). is_class G C" - app :: "instr \ jvm_prog \ nat \ ty \ nat \ exception_table \ state_type option \ bool" +definition app :: "instr \ jvm_prog \ nat \ ty \ nat \ exception_table \ state_type option \ bool" where "app i G maxs rT pc et s == case s of None \ True | Some t \ app' (i,G,pc,maxs,rT,t) \ xcpt_app i G pc et" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/JType.thy --- a/src/HOL/MicroJava/BV/JType.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/BV/JType.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,19 +9,17 @@ imports "../DFA/Semilattices" "../J/WellForm" begin -constdefs - super :: "'a prog \ cname \ cname" +definition super :: "'a prog \ cname \ cname" where "super G C == fst (the (class G C))" lemma superI: "G \ C \C1 D \ super G C = D" by (unfold super_def) (auto dest: subcls1D) -constdefs - is_ref :: "ty \ bool" +definition is_ref :: "ty \ bool" where "is_ref T == case T of PrimT t \ False | RefT r \ True" - sup :: "'c prog \ ty \ ty \ ty err" +definition sup :: "'c prog \ ty \ ty \ ty err" where "sup G T1 T2 == case T1 of PrimT P1 \ (case T2 of PrimT P2 \ (if P1 = P2 then OK (PrimT P1) else Err) | RefT R \ Err) @@ -30,17 +28,16 @@ | ClassT C \ (case R2 of NullT \ OK (Class C) | ClassT D \ OK (Class (exec_lub (subcls1 G) (super G) C D)))))" - subtype :: "'c prog \ ty \ ty \ bool" +definition subtype :: "'c prog \ ty \ ty \ bool" where "subtype G T1 T2 == G \ T1 \ T2" - is_ty :: "'c prog \ ty \ bool" +definition is_ty :: "'c prog \ ty \ bool" where "is_ty G T == case T of PrimT P \ True | RefT R \ (case R of NullT \ True | ClassT C \ (C, Object) \ (subcls1 G)^*)" abbreviation "types G == Collect (is_type G)" -constdefs - esl :: "'c prog \ ty esl" +definition esl :: "'c prog \ ty esl" where "esl G == (types G, subtype G, sup G)" lemma PrimT_PrimT: "(G \ xb \ PrimT p) = (xb = PrimT p)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/JVM.thy --- a/src/HOL/MicroJava/BV/JVM.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/BV/JVM.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,14 +9,13 @@ imports Typing_Framework_JVM begin -constdefs - kiljvm :: "jvm_prog \ nat \ nat \ ty \ exception_table \ - instr list \ JVMType.state list \ JVMType.state list" +definition kiljvm :: "jvm_prog \ nat \ nat \ ty \ exception_table \ + instr list \ JVMType.state list \ JVMType.state list" where "kiljvm G maxs maxr rT et bs == kildall (JVMType.le G maxs maxr) (JVMType.sup G maxs maxr) (exec G maxs rT et bs)" - wt_kil :: "jvm_prog \ cname \ ty list \ ty \ nat \ nat \ - exception_table \ instr list \ bool" +definition wt_kil :: "jvm_prog \ cname \ ty list \ ty \ nat \ nat \ + exception_table \ instr list \ bool" where "wt_kil G C pTs rT mxs mxl et ins == check_bounded ins et \ 0 < size ins \ (let first = Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)); @@ -24,11 +23,10 @@ result = kiljvm G mxs (1+size pTs+mxl) rT et ins start in \n < size ins. result!n \ Err)" - wt_jvm_prog_kildall :: "jvm_prog \ bool" +definition wt_jvm_prog_kildall :: "jvm_prog \ bool" where "wt_jvm_prog_kildall G == wf_prog (\G C (sig,rT,(maxs,maxl,b,et)). wt_kil G C (snd sig) rT maxs maxl et b) G" - theorem is_bcv_kiljvm: "\ wf_prog wf_mb G; bounded (exec G maxs rT et bs) (size bs) \ \ is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/JVMType.thy --- a/src/HOL/MicroJava/BV/JVMType.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/BV/JVMType.thy Mon Mar 01 13:42:31 2010 +0100 @@ -19,43 +19,39 @@ prog_type = "cname \ class_type" -constdefs - stk_esl :: "'c prog \ nat \ ty list esl" +definition stk_esl :: "'c prog \ nat \ ty list esl" where "stk_esl S maxs == upto_esl maxs (JType.esl S)" - reg_sl :: "'c prog \ nat \ ty err list sl" +definition reg_sl :: "'c prog \ nat \ ty err list sl" where "reg_sl S maxr == Listn.sl maxr (Err.sl (JType.esl S))" - sl :: "'c prog \ nat \ nat \ state sl" +definition sl :: "'c prog \ nat \ nat \ state sl" where "sl S maxs maxr == Err.sl(Opt.esl(Product.esl (stk_esl S maxs) (Err.esl(reg_sl S maxr))))" -constdefs - states :: "'c prog \ nat \ nat \ state set" +definition states :: "'c prog \ nat \ nat \ state set" where "states S maxs maxr == fst(sl S maxs maxr)" - le :: "'c prog \ nat \ nat \ state ord" +definition le :: "'c prog \ nat \ nat \ state ord" where "le S maxs maxr == fst(snd(sl S maxs maxr))" - sup :: "'c prog \ nat \ nat \ state binop" +definition sup :: "'c prog \ nat \ nat \ state binop" where "sup S maxs maxr == snd(snd(sl S maxs maxr))" - -constdefs - sup_ty_opt :: "['code prog,ty err,ty err] \ bool" - ("_ |- _ <=o _" [71,71] 70) +definition sup_ty_opt :: "['code prog,ty err,ty err] \ bool" + ("_ |- _ <=o _" [71,71] 70) where "sup_ty_opt G == Err.le (subtype G)" - sup_loc :: "['code prog,locvars_type,locvars_type] \ bool" - ("_ |- _ <=l _" [71,71] 70) +definition sup_loc :: "['code prog,locvars_type,locvars_type] \ bool" + ("_ |- _ <=l _" [71,71] 70) where "sup_loc G == Listn.le (sup_ty_opt G)" - sup_state :: "['code prog,state_type,state_type] \ bool" - ("_ |- _ <=s _" [71,71] 70) +definition sup_state :: "['code prog,state_type,state_type] \ bool" + ("_ |- _ <=s _" [71,71] 70) where "sup_state G == Product.le (Listn.le (subtype G)) (sup_loc G)" - sup_state_opt :: "['code prog,state_type option,state_type option] \ bool" - ("_ |- _ <=' _" [71,71] 70) +definition sup_state_opt :: "['code prog,state_type option,state_type option] \ bool" + ("_ |- _ <=' _" [71,71] 70) where "sup_state_opt G == Opt.le (sup_state G)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/LBVJVM.thy --- a/src/HOL/MicroJava/BV/LBVJVM.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/BV/LBVJVM.thy Mon Mar 01 13:42:31 2010 +0100 @@ -11,18 +11,17 @@ types prog_cert = "cname \ sig \ JVMType.state list" -constdefs - check_cert :: "jvm_prog \ nat \ nat \ nat \ JVMType.state list \ bool" +definition check_cert :: "jvm_prog \ nat \ nat \ nat \ JVMType.state list \ bool" where "check_cert G mxs mxr n cert \ check_types G mxs mxr cert \ length cert = n+1 \ (\i Err) \ cert!n = OK None" - lbvjvm :: "jvm_prog \ nat \ nat \ ty \ exception_table \ - JVMType.state list \ instr list \ JVMType.state \ JVMType.state" +definition lbvjvm :: "jvm_prog \ nat \ nat \ ty \ exception_table \ + JVMType.state list \ instr list \ JVMType.state \ JVMType.state" where "lbvjvm G maxs maxr rT et cert bs \ wtl_inst_list bs cert (JVMType.sup G maxs maxr) (JVMType.le G maxs maxr) Err (OK None) (exec G maxs rT et bs) 0" - wt_lbv :: "jvm_prog \ cname \ ty list \ ty \ nat \ nat \ - exception_table \ JVMType.state list \ instr list \ bool" +definition wt_lbv :: "jvm_prog \ cname \ ty list \ ty \ nat \ nat \ + exception_table \ JVMType.state list \ instr list \ bool" where "wt_lbv G C pTs rT mxs mxl et cert ins \ check_bounded ins et \ check_cert G mxs (1+size pTs+mxl) (length ins) cert \ @@ -31,15 +30,15 @@ result = lbvjvm G mxs (1+size pTs+mxl) rT et cert ins (OK start) in result \ Err)" - wt_jvm_prog_lbv :: "jvm_prog \ prog_cert \ bool" +definition wt_jvm_prog_lbv :: "jvm_prog \ prog_cert \ bool" where "wt_jvm_prog_lbv G cert \ wf_prog (\G C (sig,rT,(maxs,maxl,b,et)). wt_lbv G C (snd sig) rT maxs maxl et (cert C sig) b) G" - mk_cert :: "jvm_prog \ nat \ ty \ exception_table \ instr list - \ method_type \ JVMType.state list" +definition mk_cert :: "jvm_prog \ nat \ ty \ exception_table \ instr list + \ method_type \ JVMType.state list" where "mk_cert G maxs rT et bs phi \ make_cert (exec G maxs rT et bs) (map OK phi) (OK None)" - prg_cert :: "jvm_prog \ prog_type \ prog_cert" +definition prg_cert :: "jvm_prog \ prog_type \ prog_cert" where "prg_cert G phi C sig \ let (C,rT,(maxs,maxl,ins,et)) = the (method (G,C) sig) in mk_cert G maxs rT et ins (phi C sig)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/BV/Typing_Framework_JVM.thy --- a/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,13 +9,11 @@ imports "../DFA/Abstract_BV" JVMType EffectMono BVSpec begin -constdefs - exec :: "jvm_prog \ nat \ ty \ exception_table \ instr list \ JVMType.state step_type" +definition exec :: "jvm_prog \ nat \ ty \ exception_table \ instr list \ JVMType.state step_type" where "exec G maxs rT et bs == err_step (size bs) (\pc. app (bs!pc) G maxs rT pc et) (\pc. eff (bs!pc) G pc et)" -constdefs - opt_states :: "'c prog \ nat \ nat \ (ty list \ ty err list) option set" +definition opt_states :: "'c prog \ nat \ nat \ (ty list \ ty err list) option set" where "opt_states G maxs maxr \ opt (\{list n (types G) |n. n \ maxs} \ list maxr (err (types G)))" @@ -26,8 +24,7 @@ "list_all'_rec P n [] = True" "list_all'_rec P n (x#xs) = (P x n \ list_all'_rec P (Suc n) xs)" -constdefs - list_all' :: "('a \ nat \ bool) \ 'a list \ bool" +definition list_all' :: "('a \ nat \ bool) \ 'a list \ bool" where "list_all' P xs \ list_all'_rec P 0 xs" lemma list_all'_rec: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/Comp/CorrComp.thy --- a/src/HOL/MicroJava/Comp/CorrComp.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/Comp/CorrComp.thy Mon Mar 01 13:42:31 2010 +0100 @@ -81,13 +81,12 @@ (***********************************************************************) -constdefs - progression :: "jvm_prog \ cname \ sig \ +definition progression :: "jvm_prog \ cname \ sig \ aheap \ opstack \ locvars \ bytecode \ aheap \ opstack \ locvars \ bool" - ("{_,_,_} \ {_, _, _} >- _ \ {_, _, _}" [61,61,61,61,61,61,90,61,61,61]60) + ("{_,_,_} \ {_, _, _} >- _ \ {_, _, _}" [61,61,61,61,61,61,90,61,61,61]60) where "{G,C,S} \ {hp0, os0, lvars0} >- instrs \ {hp1, os1, lvars1} == \ pre post frs. (gis (gmb G C S) = pre @ instrs @ post) \ @@ -161,10 +160,9 @@ done (*****) -constdefs - jump_fwd :: "jvm_prog \ cname \ sig \ +definition jump_fwd :: "jvm_prog \ cname \ sig \ aheap \ locvars \ opstack \ opstack \ - instr \ bytecode \ bool" + instr \ bytecode \ bool" where "jump_fwd G C S hp lvars os0 os1 instr instrs == \ pre post frs. (gis (gmb G C S) = pre @ instr # instrs @ post) \ @@ -216,10 +214,9 @@ (* note: instrs and instr reversed wrt. jump_fwd *) -constdefs - jump_bwd :: "jvm_prog \ cname \ sig \ +definition jump_bwd :: "jvm_prog \ cname \ sig \ aheap \ locvars \ opstack \ opstack \ - bytecode \ instr \ bool" + bytecode \ instr \ bool" where "jump_bwd G C S hp lvars os0 os1 instrs instr == \ pre post frs. (gis (gmb G C S) = pre @ instrs @ instr # post) \ @@ -258,14 +255,14 @@ (**********************************************************************) (* class C with signature S is defined in program G *) -constdefs class_sig_defined :: "'c prog \ cname \ sig \ bool" +definition class_sig_defined :: "'c prog \ cname \ sig \ bool" where "class_sig_defined G C S == is_class G C \ (\ D rT mb. (method (G, C) S = Some (D, rT, mb)))" (* The environment of a java method body (characterized by class and signature) *) -constdefs env_of_jmb :: "java_mb prog \ cname \ sig \ java_mb env" +definition env_of_jmb :: "java_mb prog \ cname \ sig \ java_mb env" where "env_of_jmb G C S == (let (mn,pTs) = S; (D,rT,(pns,lvars,blk,res)) = the(method (G, C) S) in @@ -331,11 +328,13 @@ (**********************************************************************) -constdefs wtpd_expr :: "java_mb env \ expr \ bool" +definition wtpd_expr :: "java_mb env \ expr \ bool" where "wtpd_expr E e == (\ T. E\e :: T)" - wtpd_exprs :: "java_mb env \ (expr list) \ bool" + +definition wtpd_exprs :: "java_mb env \ (expr list) \ bool" where "wtpd_exprs E e == (\ T. E\e [::] T)" - wtpd_stmt :: "java_mb env \ stmt \ bool" + +definition wtpd_stmt :: "java_mb env \ stmt \ bool" where "wtpd_stmt E c == (E\c \)" lemma wtpd_expr_newc: "wtpd_expr E (NewC C) \ is_class (prg E) C" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/Comp/CorrCompTp.thy --- a/src/HOL/MicroJava/Comp/CorrCompTp.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy Mon Mar 01 13:42:31 2010 +0100 @@ -12,13 +12,13 @@ (**********************************************************************) -constdefs - inited_LT :: "[cname, ty list, (vname \ ty) list] \ locvars_type" +definition inited_LT :: "[cname, ty list, (vname \ ty) list] \ locvars_type" where "inited_LT C pTs lvars == (OK (Class C))#((map OK pTs))@(map (Fun.comp OK snd) lvars)" - is_inited_LT :: "[cname, ty list, (vname \ ty) list, locvars_type] \ bool" + +definition is_inited_LT :: "[cname, ty list, (vname \ ty) list, locvars_type] \ bool" where "is_inited_LT C pTs lvars LT == (LT = (inited_LT C pTs lvars))" - local_env :: "[java_mb prog, cname, sig, vname list,(vname \ ty) list] \ java_mb env" +definition local_env :: "[java_mb prog, cname, sig, vname list,(vname \ ty) list] \ java_mb env" where "local_env G C S pns lvars == let (mn, pTs) = S in (G,map_of lvars(pns[\]pTs)(This\Class C))" @@ -536,13 +536,12 @@ (**********************************************************************) -constdefs - offset_xcentry :: "[nat, exception_entry] \ exception_entry" +definition offset_xcentry :: "[nat, exception_entry] \ exception_entry" where "offset_xcentry == \ n (start_pc, end_pc, handler_pc, catch_type). (start_pc + n, end_pc + n, handler_pc + n, catch_type)" - offset_xctable :: "[nat, exception_table] \ exception_table" +definition offset_xctable :: "[nat, exception_table] \ exception_table" where "offset_xctable n == (map (offset_xcentry n))" lemma match_xcentry_offset [simp]: " @@ -682,12 +681,11 @@ (**********************************************************************) -constdefs - start_sttp_resp_cons :: "[state_type \ method_type \ state_type] \ bool" +definition start_sttp_resp_cons :: "[state_type \ method_type \ state_type] \ bool" where "start_sttp_resp_cons f == (\ sttp. let (mt', sttp') = (f sttp) in (\mt'_rest. mt' = Some sttp # mt'_rest))" - start_sttp_resp :: "[state_type \ method_type \ state_type] \ bool" +definition start_sttp_resp :: "[state_type \ method_type \ state_type] \ bool" where "start_sttp_resp f == (f = comb_nil) \ (start_sttp_resp_cons f)" lemma start_sttp_resp_comb_nil [simp]: "start_sttp_resp comb_nil" @@ -887,10 +885,9 @@ (* ******************************************************************* *) -constdefs - bc_mt_corresp :: " +definition bc_mt_corresp :: " [bytecode, state_type \ method_type \ state_type, state_type, jvm_prog, ty, nat, p_count] - \ bool" + \ bool" where "bc_mt_corresp bc f sttp0 cG rT mxr idx == let (mt, sttp) = f sttp0 in @@ -993,8 +990,7 @@ (* ********************************************************************** *) -constdefs - mt_sttp_flatten :: "method_type \ state_type \ method_type" +definition mt_sttp_flatten :: "method_type \ state_type \ method_type" where "mt_sttp_flatten mt_sttp == (mt_of mt_sttp) @ [Some (sttp_of mt_sttp)]" @@ -1473,8 +1469,7 @@ (* ******************** *) -constdefs - contracting :: "(state_type \ method_type \ state_type) \ bool" +definition contracting :: "(state_type \ method_type \ state_type) \ bool" where "contracting f == (\ ST LT. let (ST', LT') = sttp_of (f (ST, LT)) in (length ST' \ length ST \ set ST' \ set ST \ diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/Comp/DefsComp.thy --- a/src/HOL/MicroJava/Comp/DefsComp.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/Comp/DefsComp.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/Comp/DefsComp.thy - ID: $Id$ Author: Martin Strecker *) @@ -10,50 +9,65 @@ begin -constdefs - method_rT :: "cname \ ty \ 'c \ ty" +definition method_rT :: "cname \ ty \ 'c \ ty" where "method_rT mtd == (fst (snd mtd))" -constdefs (* g = get *) - gx :: "xstate \ val option" "gx \ fst" - gs :: "xstate \ state" "gs \ snd" - gh :: "xstate \ aheap" "gh \ fst\snd" - gl :: "xstate \ State.locals" "gl \ snd\snd" +definition + gx :: "xstate \ val option" where "gx \ fst" + +definition + gs :: "xstate \ state" where "gs \ snd" + +definition + gh :: "xstate \ aheap" where "gh \ fst\snd" +definition + gl :: "xstate \ State.locals" where "gl \ snd\snd" + +definition gmb :: "'a prog \ cname \ sig \ 'a" - "gmb G cn si \ snd(snd(the(method (G,cn) si)))" + where "gmb G cn si \ snd(snd(the(method (G,cn) si)))" + +definition gis :: "jvm_method \ bytecode" - "gis \ fst \ snd \ snd" + where "gis \ fst \ snd \ snd" (* jmb = aus einem JavaMaethodBody *) - gjmb_pns :: "java_mb \ vname list" "gjmb_pns \ fst" - gjmb_lvs :: "java_mb \ (vname\ty)list" "gjmb_lvs \ fst\snd" - gjmb_blk :: "java_mb \ stmt" "gjmb_blk \ fst\snd\snd" - gjmb_res :: "java_mb \ expr" "gjmb_res \ snd\snd\snd" +definition + gjmb_pns :: "java_mb \ vname list" where "gjmb_pns \ fst" + +definition + gjmb_lvs :: "java_mb \ (vname\ty)list" where "gjmb_lvs \ fst\snd" + +definition + gjmb_blk :: "java_mb \ stmt" where "gjmb_blk \ fst\snd\snd" + +definition + gjmb_res :: "java_mb \ expr" where "gjmb_res \ snd\snd\snd" + +definition gjmb_plns :: "java_mb \ vname list" - "gjmb_plns \ \jmb. gjmb_pns jmb @ map fst (gjmb_lvs jmb)" + where "gjmb_plns \ \jmb. gjmb_pns jmb @ map fst (gjmb_lvs jmb)" +definition glvs :: "java_mb \ State.locals \ locvars" - "glvs jmb loc \ map (the\loc) (gjmb_plns jmb)" + where "glvs jmb loc \ map (the\loc) (gjmb_plns jmb)" lemmas gdefs = gx_def gh_def gl_def gmb_def gis_def glvs_def lemmas gjmbdefs = gjmb_pns_def gjmb_lvs_def gjmb_blk_def gjmb_res_def gjmb_plns_def lemmas galldefs = gdefs gjmbdefs - - -constdefs - locvars_locals :: "java_mb prog \ cname \ sig \ State.locals \ locvars" +definition locvars_locals :: "java_mb prog \ cname \ sig \ State.locals \ locvars" where "locvars_locals G C S lvs == the (lvs This) # glvs (gmb G C S) lvs" - locals_locvars :: "java_mb prog \ cname \ sig \ locvars \ State.locals" +definition locals_locvars :: "java_mb prog \ cname \ sig \ locvars \ State.locals" where "locals_locvars G C S lvs == empty ((gjmb_plns (gmb G C S))[\](tl lvs)) (This\(hd lvs))" - locvars_xstate :: "java_mb prog \ cname \ sig \ xstate \ locvars" +definition locvars_xstate :: "java_mb prog \ cname \ sig \ xstate \ locvars" where "locvars_xstate G C S xs == locvars_locals G C S (gl xs)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/Comp/Index.thy --- a/src/HOL/MicroJava/Comp/Index.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/Comp/Index.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,8 +9,7 @@ begin (*indexing a variable name among all variable declarations in a method body*) -constdefs - index :: "java_mb => vname => nat" +definition index :: "java_mb => vname => nat" where "index == \ (pn,lv,blk,res) v. if v = This then 0 @@ -99,8 +98,7 @@ (* The following def should replace the conditions in WellType.thy / wf_java_mdecl *) -constdefs - disjoint_varnames :: "[vname list, (vname \ ty) list] \ bool" +definition disjoint_varnames :: "[vname list, (vname \ ty) list] \ bool" where (* This corresponds to the original def in wf_java_mdecl: "disjoint_varnames pns lvars \ nodups pns \ unique lvars \ This \ set pns \ This \ set (map fst lvars) \ diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/Comp/TranslComp.thy --- a/src/HOL/MicroJava/Comp/TranslComp.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/Comp/TranslComp.thy Mon Mar 01 13:42:31 2010 +0100 @@ -98,17 +98,16 @@ (*compiling methods, classes and programs*) (*initialising a single variable*) -constdefs - load_default_val :: "ty => instr" +definition load_default_val :: "ty => instr" where "load_default_val ty == LitPush (default_val ty)" - compInit :: "java_mb => (vname * ty) => instr list" +definition compInit :: "java_mb => (vname * ty) => instr list" where "compInit jmb == \ (vn,ty). [load_default_val ty, Store (index jmb vn)]" - compInitLvars :: "[java_mb, (vname \ ty) list] \ bytecode" +definition compInitLvars :: "[java_mb, (vname \ ty) list] \ bytecode" where "compInitLvars jmb lvars == concat (map (compInit jmb) lvars)" - compMethod :: "java_mb prog \ cname \ java_mb mdecl \ jvm_method mdecl" +definition compMethod :: "java_mb prog \ cname \ java_mb mdecl \ jvm_method mdecl" where "compMethod G C jmdl == let (sig, rT, jmb) = jmdl; (pns,lvars,blk,res) = jmb; mt = (compTpMethod G C jmdl); @@ -117,10 +116,10 @@ [Return] in (sig, rT, max_ssize mt, length lvars, bc, [])" - compClass :: "java_mb prog => java_mb cdecl=> jvm_method cdecl" +definition compClass :: "java_mb prog => java_mb cdecl=> jvm_method cdecl" where "compClass G == \ (C,cno,fdls,jmdls). (C,cno,fdls, map (compMethod G C) jmdls)" - comp :: "java_mb prog => jvm_prog" +definition comp :: "java_mb prog => jvm_prog" where "comp G == map (compClass G) G" end diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/Comp/TranslCompTp.thy --- a/src/HOL/MicroJava/Comp/TranslCompTp.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/Comp/TranslCompTp.thy Mon Mar 01 13:42:31 2010 +0100 @@ -6,17 +6,14 @@ imports Index "../BV/JVMType" begin - - (**********************************************************************) - -constdefs - comb :: "['a \ 'b list \ 'c, 'c \ 'b list \ 'd, 'a] \ 'b list \ 'd" +definition comb :: "['a \ 'b list \ 'c, 'c \ 'b list \ 'd, 'a] \ 'b list \ 'd" where "comb == (\ f1 f2 x0. let (xs1, x1) = f1 x0; (xs2, x2) = f2 x1 in (xs1 @ xs2, x2))" - comb_nil :: "'a \ 'b list \ 'a" + +definition comb_nil :: "'a \ 'b list \ 'a" where "comb_nil a == ([], a)" notation (xsymbols) @@ -58,23 +55,26 @@ compTpStmt :: "java_mb \ java_mb prog \ stmt \ state_type \ method_type \ state_type" - -constdefs - nochangeST :: "state_type \ method_type \ state_type" +definition nochangeST :: "state_type \ method_type \ state_type" where "nochangeST sttp == ([Some sttp], sttp)" - pushST :: "[ty list, state_type] \ method_type \ state_type" + +definition pushST :: "[ty list, state_type] \ method_type \ state_type" where "pushST tps == (\ (ST, LT). ([Some (ST, LT)], (tps @ ST, LT)))" - dupST :: "state_type \ method_type \ state_type" + +definition dupST :: "state_type \ method_type \ state_type" where "dupST == (\ (ST, LT). ([Some (ST, LT)], (hd ST # ST, LT)))" - dup_x1ST :: "state_type \ method_type \ state_type" + +definition dup_x1ST :: "state_type \ method_type \ state_type" where "dup_x1ST == (\ (ST, LT). ([Some (ST, LT)], (hd ST # hd (tl ST) # hd ST # (tl (tl ST)), LT)))" - popST :: "[nat, state_type] \ method_type \ state_type" + +definition popST :: "[nat, state_type] \ method_type \ state_type" where "popST n == (\ (ST, LT). ([Some (ST, LT)], (drop n ST, LT)))" - replST :: "[nat, ty, state_type] \ method_type \ state_type" + +definition replST :: "[nat, ty, state_type] \ method_type \ state_type" where "replST n tp == (\ (ST, LT). ([Some (ST, LT)], (tp # (drop n ST), LT)))" - storeST :: "[nat, ty, state_type] \ method_type \ state_type" +definition storeST :: "[nat, ty, state_type] \ method_type \ state_type" where "storeST i tp == (\ (ST, LT). ([Some (ST, LT)], (tl ST, LT [i:= OK tp])))" @@ -137,9 +137,8 @@ (pushST [PrimT Boolean]) \ (compTpExpr jmb G e) \ popST 2 \ (compTpStmt jmb G c) \ nochangeST" -constdefs - compTpInit :: "java_mb \ (vname * ty) - \ state_type \ method_type \ state_type" +definition compTpInit :: "java_mb \ (vname * ty) + \ state_type \ method_type \ state_type" where "compTpInit jmb == (\ (vn,ty). (pushST [ty]) \ (storeST (index jmb vn) ty))" consts @@ -150,14 +149,13 @@ "compTpInitLvars jmb [] = comb_nil" "compTpInitLvars jmb (lv#lvars) = (compTpInit jmb lv) \ (compTpInitLvars jmb lvars)" -constdefs - start_ST :: "opstack_type" +definition start_ST :: "opstack_type" where "start_ST == []" - start_LT :: "cname \ ty list \ nat \ locvars_type" +definition start_LT :: "cname \ ty list \ nat \ locvars_type" where "start_LT C pTs n == (OK (Class C))#((map OK pTs))@(replicate n Err)" - compTpMethod :: "[java_mb prog, cname, java_mb mdecl] \ method_type" +definition compTpMethod :: "[java_mb prog, cname, java_mb mdecl] \ method_type" where "compTpMethod G C == \ ((mn,pTs),rT, jmb). let (pns,lvars,blk,res) = jmb in (mt_of @@ -167,7 +165,7 @@ nochangeST) (start_ST, start_LT C pTs (length lvars))))" - compTp :: "java_mb prog => prog_type" +definition compTp :: "java_mb prog => prog_type" where "compTp G C sig == let (D, rT, jmb) = (the (method (G, C) sig)) in compTpMethod G C (sig, rT, jmb)" @@ -176,15 +174,13 @@ (**********************************************************************) (* Computing the maximum stack size from the method_type *) -constdefs - ssize_sto :: "(state_type option) \ nat" +definition ssize_sto :: "(state_type option) \ nat" where "ssize_sto sto == case sto of None \ 0 | (Some (ST, LT)) \ length ST" - max_of_list :: "nat list \ nat" +definition max_of_list :: "nat list \ nat" where "max_of_list xs == foldr max xs 0" - max_ssize :: "method_type \ nat" +definition max_ssize :: "method_type \ nat" where "max_ssize mt == max_of_list (map ssize_sto mt)" - end diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/Comp/TypeInf.thy --- a/src/HOL/MicroJava/Comp/TypeInf.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/Comp/TypeInf.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/Comp/TypeInf.thy - ID: $Id$ Author: Martin Strecker *) @@ -169,10 +168,10 @@ -constdefs - inferred_tp :: "[java_mb env, expr] \ ty" +definition inferred_tp :: "[java_mb env, expr] \ ty" where "inferred_tp E e == (SOME T. E\e :: T)" - inferred_tps :: "[java_mb env, expr list] \ ty list" + +definition inferred_tps :: "[java_mb env, expr list] \ ty list" where "inferred_tps E es == (SOME Ts. E\es [::] Ts)" (* get inferred type(s) for well-typed term *) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/DFA/Err.thy --- a/src/HOL/MicroJava/DFA/Err.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/DFA/Err.thy Mon Mar 01 13:42:31 2010 +0100 @@ -14,35 +14,32 @@ types 'a ebinop = "'a \ 'a \ 'a err" 'a esl = "'a set * 'a ord * 'a ebinop" -consts - ok_val :: "'a err \ 'a" -primrec +primrec ok_val :: "'a err \ 'a" where "ok_val (OK x) = x" -constdefs - lift :: "('a \ 'b err) \ ('a err \ 'b err)" +definition lift :: "('a \ 'b err) \ ('a err \ 'b err)" where "lift f e == case e of Err \ Err | OK x \ f x" - lift2 :: "('a \ 'b \ 'c err) \ 'a err \ 'b err \ 'c err" +definition lift2 :: "('a \ 'b \ 'c err) \ 'a err \ 'b err \ 'c err" where "lift2 f e1 e2 == case e1 of Err \ Err | OK x \ (case e2 of Err \ Err | OK y \ f x y)" - le :: "'a ord \ 'a err ord" +definition le :: "'a ord \ 'a err ord" where "le r e1 e2 == case e2 of Err \ True | OK y \ (case e1 of Err \ False | OK x \ x <=_r y)" - sup :: "('a \ 'b \ 'c) \ ('a err \ 'b err \ 'c err)" +definition sup :: "('a \ 'b \ 'c) \ ('a err \ 'b err \ 'c err)" where "sup f == lift2(%x y. OK(x +_f y))" - err :: "'a set \ 'a err set" +definition err :: "'a set \ 'a err set" where "err A == insert Err {x . ? y:A. x = OK y}" - esl :: "'a sl \ 'a esl" +definition esl :: "'a sl \ 'a esl" where "esl == %(A,r,f). (A,r, %x y. OK(f x y))" - sl :: "'a esl \ 'a err sl" +definition sl :: "'a esl \ 'a err sl" where "sl == %(A,r,f). (err A, le r, lift2 f)" abbreviation diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/DFA/Kildall.thy --- a/src/HOL/MicroJava/DFA/Kildall.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/DFA/Kildall.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,37 +9,28 @@ imports SemilatAlg While_Combinator begin - -consts - iter :: "'s binop \ 's step_type \ - 's list \ nat set \ 's list \ nat set" - propa :: "'s binop \ (nat \ 's) list \ 's list \ nat set \ 's list * nat set" +primrec propa :: "'s binop \ (nat \ 's) list \ 's list \ nat set \ 's list * nat set" where + "propa f [] ss w = (ss,w)" +| "propa f (q'#qs) ss w = (let (q,t) = q'; + u = t +_f ss!q; + w' = (if u = ss!q then w else insert q w) + in propa f qs (ss[q := u]) w')" -primrec -"propa f [] ss w = (ss,w)" -"propa f (q'#qs) ss w = (let (q,t) = q'; - u = t +_f ss!q; - w' = (if u = ss!q then w else insert q w) - in propa f qs (ss[q := u]) w')" - -defs iter_def: -"iter f step ss w == - while (%(ss,w). w \ {}) +definition iter :: "'s binop \ 's step_type \ 's list \ nat set \ 's list \ nat set" where + "iter f step ss w == while (%(ss,w). w \ {}) (%(ss,w). let p = SOME p. p \ w in propa f (step p (ss!p)) ss (w-{p})) (ss,w)" -constdefs - unstables :: "'s ord \ 's step_type \ 's list \ nat set" +definition unstables :: "'s ord \ 's step_type \ 's list \ nat set" where "unstables r step ss == {p. p < size ss \ \stable r step ss p}" - kildall :: "'s ord \ 's binop \ 's step_type \ 's list \ 's list" +definition kildall :: "'s ord \ 's binop \ 's step_type \ 's list \ 's list" where "kildall r f step ss == fst(iter f step ss (unstables r step ss))" -consts merges :: "'s binop \ (nat \ 's) list \ 's list \ 's list" -primrec -"merges f [] ss = ss" -"merges f (p'#ps) ss = (let (p,s) = p' in merges f ps (ss[p := s +_f ss!p]))" +primrec merges :: "'s binop \ (nat \ 's) list \ 's list \ 's list" where + "merges f [] ss = ss" +| "merges f (p'#ps) ss = (let (p,s) = p' in merges f ps (ss[p := s +_f ss!p]))" lemmas [simp] = Let_def Semilat.le_iff_plus_unchanged [OF Semilat.intro, symmetric] diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/DFA/LBVComplete.thy --- a/src/HOL/MicroJava/DFA/LBVComplete.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/DFA/LBVComplete.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,12 +9,11 @@ imports LBVSpec Typing_Framework begin -constdefs - is_target :: "['s step_type, 's list, nat] \ bool" +definition is_target :: "['s step_type, 's list, nat] \ bool" where "is_target step phi pc' \ \pc s'. pc' \ pc+1 \ pc < length phi \ (pc',s') \ set (step pc (phi!pc))" - make_cert :: "['s step_type, 's list, 's] \ 's certificate" +definition make_cert :: "['s step_type, 's list, 's] \ 's certificate" where "make_cert step phi B \ map (\pc. if is_target step phi pc then phi!pc else B) [0.. 's binop \ 's ord \ 's \ nat \ (nat \ 's) list \ 's \ 's" -primrec -"merge cert f r T pc [] x = x" -"merge cert f r T pc (s#ss) x = merge cert f r T pc ss (let (pc',s') = s in +primrec merge :: "'s certificate \ 's binop \ 's ord \ 's \ nat \ (nat \ 's) list \ 's \ 's" where + "merge cert f r T pc [] x = x" +| "merge cert f r T pc (s#ss) x = merge cert f r T pc ss (let (pc',s') = s in if pc'=pc+1 then s' +_f x else if s' <=_r (cert!pc') then x else T)" -constdefs -wtl_inst :: "'s certificate \ 's binop \ 's ord \ 's \ - 's step_type \ nat \ 's \ 's" +definition wtl_inst :: "'s certificate \ 's binop \ 's ord \ 's \ + 's step_type \ nat \ 's \ 's" where "wtl_inst cert f r T step pc s \ merge cert f r T pc (step pc s) (cert!(pc+1))" -wtl_cert :: "'s certificate \ 's binop \ 's ord \ 's \ 's \ - 's step_type \ nat \ 's \ 's" +definition wtl_cert :: "'s certificate \ 's binop \ 's ord \ 's \ 's \ + 's step_type \ nat \ 's \ 's" where "wtl_cert cert f r T B step pc s \ if cert!pc = B then wtl_inst cert f r T step pc s else if s <=_r (cert!pc) then wtl_inst cert f r T step pc (cert!pc) else T" -consts -wtl_inst_list :: "'a list \ 's certificate \ 's binop \ 's ord \ 's \ 's \ - 's step_type \ nat \ 's \ 's" -primrec -"wtl_inst_list [] cert f r T B step pc s = s" -"wtl_inst_list (i#is) cert f r T B step pc s = +primrec wtl_inst_list :: "'a list \ 's certificate \ 's binop \ 's ord \ 's \ 's \ + 's step_type \ nat \ 's \ 's" where + "wtl_inst_list [] cert f r T B step pc s = s" +| "wtl_inst_list (i#is) cert f r T B step pc s = (let s' = wtl_cert cert f r T B step pc s in if s' = T \ s = T then T else wtl_inst_list is cert f r T B step (pc+1) s')" -constdefs - cert_ok :: "'s certificate \ nat \ 's \ 's \ 's set \ bool" +definition cert_ok :: "'s certificate \ nat \ 's \ 's \ 's set \ bool" where "cert_ok cert n T B A \ (\i < n. cert!i \ A \ cert!i \ T) \ (cert!n = B)" -constdefs - bottom :: "'a ord \ 'a \ bool" +definition bottom :: "'a ord \ 'a \ bool" where "bottom r B \ \x. B <=_r x" - locale lbv = Semilat + fixes T :: "'a" ("\") fixes B :: "'a" ("\") diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/DFA/Listn.thy --- a/src/HOL/MicroJava/DFA/Listn.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/DFA/Listn.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,12 +9,10 @@ imports Err begin -constdefs - - list :: "nat \ 'a set \ 'a list set" +definition list :: "nat \ 'a set \ 'a list set" where "list n A == {xs. length xs = n & set xs <= A}" - le :: "'a ord \ ('a list)ord" +definition le :: "'a ord \ ('a list)ord" where "le r == list_all2 (%x y. x <=_r y)" abbreviation @@ -27,8 +25,7 @@ ("(_ /<[_] _)" [50, 0, 51] 50) where "x <[r] y == x <_(le r) y" -constdefs - map2 :: "('a \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" +definition map2 :: "('a \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" where "map2 f == (%xs ys. map (split f) (zip xs ys))" abbreviation @@ -36,19 +33,17 @@ ("(_ /+[_] _)" [65, 0, 66] 65) where "x +[f] y == x +_(map2 f) y" -consts coalesce :: "'a err list \ 'a list err" -primrec -"coalesce [] = OK[]" -"coalesce (ex#exs) = Err.sup (op #) ex (coalesce exs)" +primrec coalesce :: "'a err list \ 'a list err" where + "coalesce [] = OK[]" +| "coalesce (ex#exs) = Err.sup (op #) ex (coalesce exs)" -constdefs - sl :: "nat \ 'a sl \ 'a list sl" +definition sl :: "nat \ 'a sl \ 'a list sl" where "sl n == %(A,r,f). (list n A, le r, map2 f)" - sup :: "('a \ 'b \ 'c err) \ 'a list \ 'b list \ 'c list err" +definition sup :: "('a \ 'b \ 'c err) \ 'a list \ 'b list \ 'c list err" where "sup f == %xs ys. if size xs = size ys then coalesce(xs +[f] ys) else Err" - upto_esl :: "nat \ 'a esl \ 'a list esl" +definition upto_esl :: "nat \ 'a esl \ 'a list esl" where "upto_esl m == %(A,r,f). (Union{list n A |n. n <= m}, le r, sup f)" lemmas [simp] = set_update_subsetI diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/DFA/Opt.thy --- a/src/HOL/MicroJava/DFA/Opt.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/DFA/Opt.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,21 +9,20 @@ imports Err begin -constdefs - le :: "'a ord \ 'a option ord" +definition le :: "'a ord \ 'a option ord" where "le r o1 o2 == case o2 of None \ o1=None | Some y \ (case o1 of None \ True | Some x \ x <=_r y)" - opt :: "'a set \ 'a option set" +definition opt :: "'a set \ 'a option set" where "opt A == insert None {x . ? y:A. x = Some y}" - sup :: "'a ebinop \ 'a option ebinop" +definition sup :: "'a ebinop \ 'a option ebinop" where "sup f o1 o2 == case o1 of None \ OK o2 | Some x \ (case o2 of None \ OK o1 | Some y \ (case f x y of Err \ Err | OK z \ OK (Some z)))" - esl :: "'a esl \ 'a option esl" +definition esl :: "'a esl \ 'a option esl" where "esl == %(A,r,f). (opt A, le r, sup f)" lemma unfold_le_opt: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/DFA/Product.thy --- a/src/HOL/MicroJava/DFA/Product.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/DFA/Product.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,14 +9,13 @@ imports Err begin -constdefs - le :: "'a ord \ 'b ord \ ('a * 'b) ord" +definition le :: "'a ord \ 'b ord \ ('a * 'b) ord" where "le rA rB == %(a,b) (a',b'). a <=_rA a' & b <=_rB b'" - sup :: "'a ebinop \ 'b ebinop \ ('a * 'b)ebinop" +definition sup :: "'a ebinop \ 'b ebinop \ ('a * 'b)ebinop" where "sup f g == %(a1,b1)(a2,b2). Err.sup Pair (a1 +_f a2) (b1 +_g b2)" - esl :: "'a esl \ 'b esl \ ('a * 'b ) esl" +definition esl :: "'a esl \ 'b esl \ ('a * 'b ) esl" where "esl == %(A,rA,fA) (B,rB,fB). (A <*> B, le rA rB, sup fA fB)" abbreviation diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/DFA/Semilat.thy --- a/src/HOL/MicroJava/DFA/Semilat.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/DFA/Semilat.thy Mon Mar 01 13:42:31 2010 +0100 @@ -52,36 +52,34 @@ lesssub_def: "x \\<^sub>r y \ x \\<^sub>r y \ x \ y" plussub_def: "x \\<^sub>f y \ f x y" -constdefs - ord :: "('a \ 'a) set \ 'a ord" +definition ord :: "('a \ 'a) set \ 'a ord" where "ord r \ \x y. (x,y) \ r" - order :: "'a ord \ bool" +definition order :: "'a ord \ bool" where "order r \ (\x. x \\<^sub>r x) \ (\x y. x \\<^sub>r y \ y \\<^sub>r x \ x=y) \ (\x y z. x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z)" - top :: "'a ord \ 'a \ bool" +definition top :: "'a ord \ 'a \ bool" where "top r T \ \x. x \\<^sub>r T" - acc :: "'a ord \ bool" +definition acc :: "'a ord \ bool" where "acc r \ wf {(y,x). x \\<^sub>r y}" - closed :: "'a set \ 'a binop \ bool" +definition closed :: "'a set \ 'a binop \ bool" where "closed A f \ \x\A. \y\A. x \\<^sub>f y \ A" - semilat :: "'a sl \ bool" +definition semilat :: "'a sl \ bool" where "semilat \ \(A,r,f). order r \ closed A f \ (\x\A. \y\A. x \\<^sub>r x \\<^sub>f y) \ (\x\A. \y\A. y \\<^sub>r x \\<^sub>f y) \ (\x\A. \y\A. \z\A. x \\<^sub>r z \ y \\<^sub>r z \ x \\<^sub>f y \\<^sub>r z)" - - is_ub :: "('a \ 'a) set \ 'a \ 'a \ 'a \ bool" +definition is_ub :: "('a \ 'a) set \ 'a \ 'a \ 'a \ bool" where "is_ub r x y u \ (x,u)\r \ (y,u)\r" - is_lub :: "('a \ 'a) set \ 'a \ 'a \ 'a \ bool" +definition is_lub :: "('a \ 'a) set \ 'a \ 'a \ 'a \ bool" where "is_lub r x y u \ is_ub r x y u \ (\z. is_ub r x y z \ (u,z)\r)" - some_lub :: "('a \ 'a) set \ 'a \ 'a \ 'a" +definition some_lub :: "('a \ 'a) set \ 'a \ 'a \ 'a" where "some_lub r x y \ SOME z. is_lub r x y z" locale Semilat = @@ -307,8 +305,7 @@ subsection{*An executable lub-finder*} -constdefs - exec_lub :: "('a * 'a) set \ ('a \ 'a) \ 'a binop" +definition exec_lub :: "('a * 'a) set \ ('a \ 'a) \ 'a binop" where "exec_lub r f x y \ while (\z. (x,z) \ r\<^sup>*) f y" lemma exec_lub_refl: "exec_lub r f T T = T" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/DFA/SemilatAlg.thy --- a/src/HOL/MicroJava/DFA/SemilatAlg.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/DFA/SemilatAlg.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,29 +9,24 @@ imports Typing_Framework Product begin -constdefs - lesubstep_type :: "(nat \ 's) list \ 's ord \ (nat \ 's) list \ bool" - ("(_ /<=|_| _)" [50, 0, 51] 50) +definition lesubstep_type :: "(nat \ 's) list \ 's ord \ (nat \ 's) list \ bool" + ("(_ /<=|_| _)" [50, 0, 51] 50) where "x <=|r| y \ \(p,s) \ set x. \s'. (p,s') \ set y \ s <=_r s'" -consts - plusplussub :: "'a list \ ('a \ 'a \ 'a) \ 'a \ 'a" ("(_ /++'__ _)" [65, 1000, 66] 65) -primrec +primrec plusplussub :: "'a list \ ('a \ 'a \ 'a) \ 'a \ 'a" ("(_ /++'__ _)" [65, 1000, 66] 65) where "[] ++_f y = y" - "(x#xs) ++_f y = xs ++_f (x +_f y)" +| "(x#xs) ++_f y = xs ++_f (x +_f y)" -constdefs - bounded :: "'s step_type \ nat \ bool" +definition bounded :: "'s step_type \ nat \ bool" where "bounded step n == !p nat \ 's set \ bool" +definition pres_type :: "'s step_type \ nat \ 's set \ bool" where "pres_type step n A == \s\A. \p(q,s')\set (step p s). s' \ A" - mono :: "'s ord \ 's step_type \ nat \ 's set \ bool" +definition mono :: "'s ord \ 's step_type \ nat \ 's set \ bool" where "mono r step n A == \s p t. s \ A \ p < n \ s <=_r t \ step p s <=|r| step p t" - lemma pres_typeD: "\ pres_type step n A; s\A; pset (step p s) \ \ s' \ A" by (unfold pres_type_def, blast) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/DFA/Typing_Framework.thy --- a/src/HOL/MicroJava/DFA/Typing_Framework.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/DFA/Typing_Framework.thy Mon Mar 01 13:42:31 2010 +0100 @@ -15,20 +15,19 @@ types 's step_type = "nat \ 's \ (nat \ 's) list" -constdefs - stable :: "'s ord \ 's step_type \ 's list \ nat \ bool" +definition stable :: "'s ord \ 's step_type \ 's list \ nat \ bool" where "stable r step ss p == !(q,s'):set(step p (ss!p)). s' <=_r ss!q" - stables :: "'s ord \ 's step_type \ 's list \ bool" +definition stables :: "'s ord \ 's step_type \ 's list \ bool" where "stables r step ss == !p 's \ 's step_type \ 's list \ bool" +definition wt_step :: +"'s ord \ 's \ 's step_type \ 's list \ bool" where "wt_step r T step ts == !p 's \ 's step_type - \ nat \ 's set \ ('s list \ 's list) \ bool" +definition is_bcv :: "'s ord \ 's \ 's step_type + \ nat \ 's set \ ('s list \ 's list) \ bool" where "is_bcv r T step n A bcv == !ss : list n A. (!p 's err step_type \ 's err list \ bool" +definition wt_err_step :: "'s ord \ 's err step_type \ 's err list \ bool" where "wt_err_step r step ts \ wt_step (Err.le r) Err step ts" -wt_app_eff :: "'s ord \ (nat \ 's \ bool) \ 's step_type \ 's list \ bool" +definition wt_app_eff :: "'s ord \ (nat \ 's \ bool) \ 's step_type \ 's list \ bool" where "wt_app_eff r app step ts \ \p < size ts. app p (ts!p) \ (\(q,t) \ set (step p (ts!p)). t <=_r ts!q)" -map_snd :: "('b \ 'c) \ ('a \ 'b) list \ ('a \ 'c) list" +definition map_snd :: "('b \ 'c) \ ('a \ 'b) list \ ('a \ 'c) list" where "map_snd f \ map (\(x,y). (x, f y))" -error :: "nat \ (nat \ 'a err) list" +definition error :: "nat \ (nat \ 'a err) list" where "error n \ map (\x. (x,Err)) [0.. (nat \ 's \ bool) \ 's step_type \ 's err step_type" +definition err_step :: "nat \ (nat \ 's \ bool) \ 's step_type \ 's err step_type" where "err_step n app step p t \ case t of Err \ error n | OK t' \ if app p t' then map_snd OK (step p t') else error n" -app_mono :: "'s ord \ (nat \ 's \ bool) \ nat \ 's set \ bool" +definition app_mono :: "'s ord \ (nat \ 's \ bool) \ nat \ 's set \ bool" where "app_mono r app n A \ \s p t. s \ A \ p < n \ s <=_r t \ app p t \ app p s" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/Conform.thy --- a/src/HOL/MicroJava/J/Conform.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/Conform.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/J/Conform.thy - ID: $Id$ Author: David von Oheimb Copyright 1999 Technische Universitaet Muenchen *) @@ -10,29 +9,27 @@ types 'c env' = "'c prog \ (vname \ ty)" -- "same as @{text env} of @{text WellType.thy}" -constdefs - - hext :: "aheap => aheap => bool" ("_ <=| _" [51,51] 50) +definition hext :: "aheap => aheap => bool" ("_ <=| _" [51,51] 50) where "h<=|h' == \a C fs. h a = Some(C,fs) --> (\fs'. h' a = Some(C,fs'))" - conf :: "'c prog => aheap => val => ty => bool" - ("_,_ |- _ ::<= _" [51,51,51,51] 50) +definition conf :: "'c prog => aheap => val => ty => bool" + ("_,_ |- _ ::<= _" [51,51,51,51] 50) where "G,h|-v::<=T == \T'. typeof (Option.map obj_ty o h) v = Some T' \ G\T'\T" - lconf :: "'c prog => aheap => ('a \ val) => ('a \ ty) => bool" - ("_,_ |- _ [::<=] _" [51,51,51,51] 50) +definition lconf :: "'c prog => aheap => ('a \ val) => ('a \ ty) => bool" + ("_,_ |- _ [::<=] _" [51,51,51,51] 50) where "G,h|-vs[::<=]Ts == \n T. Ts n = Some T --> (\v. vs n = Some v \ G,h|-v::<=T)" - oconf :: "'c prog => aheap => obj => bool" ("_,_ |- _ [ok]" [51,51,51] 50) +definition oconf :: "'c prog => aheap => obj => bool" ("_,_ |- _ [ok]" [51,51,51] 50) where "G,h|-obj [ok] == G,h|-snd obj[::<=]map_of (fields (G,fst obj))" - hconf :: "'c prog => aheap => bool" ("_ |-h _ [ok]" [51,51] 50) +definition hconf :: "'c prog => aheap => bool" ("_ |-h _ [ok]" [51,51] 50) where "G|-h h [ok] == \a obj. h a = Some obj --> G,h|-obj [ok]" - xconf :: "aheap \ val option \ bool" +definition xconf :: "aheap \ val option \ bool" where "xconf hp vo == preallocated hp \ (\ v. (vo = Some v) \ (\ xc. v = (Addr (XcptRef xc))))" - conforms :: "xstate => java_mb env' => bool" ("_ ::<= _" [51,51] 50) +definition conforms :: "xstate => java_mb env' => bool" ("_ ::<= _" [51,51] 50) where "s::<=E == prg E|-h heap (store s) [ok] \ prg E,heap (store s)|-locals (store s)[::<=]localT E \ xconf (heap (store s)) (abrupt s)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/Decl.thy --- a/src/HOL/MicroJava/J/Decl.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/Decl.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/J/Decl.thy - ID: $Id$ Author: David von Oheimb Copyright 1999 Technische Universitaet Muenchen *) @@ -32,11 +31,10 @@ "prog c" <= (type) "(c cdecl) list" -constdefs - "class" :: "'c prog => (cname \ 'c class)" +definition "class" :: "'c prog => (cname \ 'c class)" where "class \ map_of" - is_class :: "'c prog => cname => bool" +definition is_class :: "'c prog => cname => bool" where "is_class G C \ class G C \ None" @@ -46,10 +44,8 @@ apply (rule finite_dom_map_of) done -consts - is_type :: "'c prog => ty => bool" -primrec +primrec is_type :: "'c prog => ty => bool" where "is_type G (PrimT pt) = True" - "is_type G (RefT t) = (case t of NullT => True | ClassT C => is_class G C)" +| "is_type G (RefT t) = (case t of NullT => True | ClassT C => is_class G C)" end diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/Eval.thy --- a/src/HOL/MicroJava/J/Eval.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/Eval.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/J/Eval.thy - ID: $Id$ Author: David von Oheimb Copyright 1999 Technische Universitaet Muenchen *) @@ -11,22 +10,16 @@ -- "Auxiliary notions" -constdefs - fits :: "java_mb prog \ state \ val \ ty \ bool" ("_,_\_ fits _"[61,61,61,61]60) +definition fits :: "java_mb prog \ state \ val \ ty \ bool" ("_,_\_ fits _"[61,61,61,61]60) where "G,s\a' fits T \ case T of PrimT T' \ False | RefT T' \ a'=Null \ G\obj_ty(lookup_obj s a')\T" -constdefs - catch ::"java_mb prog \ xstate \ cname \ bool" ("_,_\catch _"[61,61,61]60) +definition catch :: "java_mb prog \ xstate \ cname \ bool" ("_,_\catch _"[61,61,61]60) where "G,s\catch C\ case abrupt s of None \ False | Some a \ G,store s\ a fits Class C" - - -constdefs - lupd :: "vname \ val \ state \ state" ("lupd'(_\_')"[10,10]1000) +definition lupd :: "vname \ val \ state \ state" ("lupd'(_\_')"[10,10]1000) where "lupd vn v \ \ (hp,loc). (hp, (loc(vn\v)))" -constdefs - new_xcpt_var :: "vname \ xstate \ xstate" +definition new_xcpt_var :: "vname \ xstate \ xstate" where "new_xcpt_var vn \ \(x,s). Norm (lupd(vn\the x) s)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/Exceptions.thy --- a/src/HOL/MicroJava/J/Exceptions.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/Exceptions.thy Mon Mar 01 13:42:31 2010 +0100 @@ -6,11 +6,10 @@ theory Exceptions imports State begin text {* a new, blank object with default values in all fields: *} -constdefs - blank :: "'c prog \ cname \ obj" +definition blank :: "'c prog \ cname \ obj" where "blank G C \ (C,init_vars (fields(G,C)))" - start_heap :: "'c prog \ aheap" +definition start_heap :: "'c prog \ aheap" where "start_heap G \ empty (XcptRef NullPointer \ blank G (Xcpt NullPointer)) (XcptRef ClassCast \ blank G (Xcpt ClassCast)) (XcptRef OutOfMemory \ blank G (Xcpt OutOfMemory))" @@ -21,8 +20,7 @@ where "cname_of hp v == fst (the (hp (the_Addr v)))" -constdefs - preallocated :: "aheap \ bool" +definition preallocated :: "aheap \ bool" where "preallocated hp \ \x. \fs. hp (XcptRef x) = Some (Xcpt x, fs)" lemma preallocatedD: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/JBasis.thy --- a/src/HOL/MicroJava/J/JBasis.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/JBasis.thy Mon Mar 01 13:42:31 2010 +0100 @@ -15,8 +15,7 @@ section "unique" -constdefs - unique :: "('a \ 'b) list => bool" +definition unique :: "('a \ 'b) list => bool" where "unique == distinct \ map fst" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/JListExample.thy --- a/src/HOL/MicroJava/J/JListExample.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/JListExample.thy Mon Mar 01 13:42:31 2010 +0100 @@ -21,29 +21,28 @@ l3_nam :: vnam l4_nam :: vnam -constdefs - val_name :: vname +definition val_name :: vname where "val_name == VName val_nam" - next_name :: vname +definition next_name :: vname where "next_name == VName next_nam" - l_name :: vname +definition l_name :: vname where "l_name == VName l_nam" - l1_name :: vname +definition l1_name :: vname where "l1_name == VName l1_nam" - l2_name :: vname +definition l2_name :: vname where "l2_name == VName l2_nam" - l3_name :: vname +definition l3_name :: vname where "l3_name == VName l3_nam" - l4_name :: vname +definition l4_name :: vname where "l4_name == VName l4_nam" - list_class :: "java_mb class" +definition list_class :: "java_mb class" where "list_class == (Object, [(val_name, PrimT Integer), (next_name, RefT (ClassT list_name))], @@ -56,7 +55,7 @@ append_name({[RefT (ClassT list_name)]}[LAcc l_name])), Lit Unit))])" - example_prg :: "java_mb prog" +definition example_prg :: "java_mb prog" where "example_prg == [ObjectC, (list_name, list_class)]" types_code diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/State.thy --- a/src/HOL/MicroJava/J/State.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/State.thy Mon Mar 01 13:42:31 2010 +0100 @@ -14,11 +14,10 @@ obj = "cname \ fields'" -- "class instance with class name and fields" -constdefs - obj_ty :: "obj => ty" +definition obj_ty :: "obj => ty" where "obj_ty obj == Class (fst obj)" - init_vars :: "('a \ ty) list => ('a \ val)" +definition init_vars :: "('a \ ty) list => ('a \ val)" where "init_vars == map_of o map (\(n,T). (n,default_val T))" types aheap = "loc \ obj" -- {* "@{text heap}" used in a translation below *} @@ -49,21 +48,19 @@ lookup_obj :: "state \ val \ obj" where "lookup_obj s a' == the (heap s (the_Addr a'))" - -constdefs - raise_if :: "bool \ xcpt \ val option \ val option" +definition raise_if :: "bool \ xcpt \ val option \ val option" where "raise_if b x xo \ if b \ (xo = None) then Some (Addr (XcptRef x)) else xo" - new_Addr :: "aheap => loc \ val option" +definition new_Addr :: "aheap => loc \ val option" where "new_Addr h \ SOME (a,x). (h a = None \ x = None) | x = Some (Addr (XcptRef OutOfMemory))" - np :: "val => val option => val option" +definition np :: "val => val option => val option" where "np v == raise_if (v = Null) NullPointer" - c_hupd :: "aheap => xstate => xstate" +definition c_hupd :: "aheap => xstate => xstate" where "c_hupd h'== \(xo,(h,l)). if xo = None then (None,(h',l)) else (xo,(h,l))" - cast_ok :: "'c prog => cname => aheap => val => bool" +definition cast_ok :: "'c prog => cname => aheap => val => bool" where "cast_ok G C h v == v = Null \ G\obj_ty (the (h (the_Addr v)))\ Class C" lemma obj_ty_def2 [simp]: "obj_ty (C,fs) = Class C" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/SystemClasses.thy --- a/src/HOL/MicroJava/J/SystemClasses.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/SystemClasses.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/J/SystemClasses.thy - ID: $Id$ Author: Gerwin Klein Copyright 2002 Technische Universitaet Muenchen *) @@ -13,20 +12,19 @@ and the system exceptions. *} -constdefs - ObjectC :: "'c cdecl" +definition ObjectC :: "'c cdecl" where "ObjectC \ (Object, (undefined,[],[]))" - NullPointerC :: "'c cdecl" +definition NullPointerC :: "'c cdecl" where "NullPointerC \ (Xcpt NullPointer, (Object,[],[]))" - ClassCastC :: "'c cdecl" +definition ClassCastC :: "'c cdecl" where "ClassCastC \ (Xcpt ClassCast, (Object,[],[]))" - OutOfMemoryC :: "'c cdecl" +definition OutOfMemoryC :: "'c cdecl" where "OutOfMemoryC \ (Xcpt OutOfMemory, (Object,[],[]))" - SystemClasses :: "'c cdecl list" +definition SystemClasses :: "'c cdecl list" where "SystemClasses \ [ObjectC, NullPointerC, ClassCastC, OutOfMemoryC]" end diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/TypeRel.thy --- a/src/HOL/MicroJava/J/TypeRel.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/TypeRel.thy Mon Mar 01 13:42:31 2010 +0100 @@ -54,9 +54,8 @@ apply auto done -constdefs - class_rec :: "'c prog \ cname \ 'a \ - (cname \ fdecl list \ 'c mdecl list \ 'a \ 'a) \ 'a" +definition class_rec :: "'c prog \ cname \ 'a \ + (cname \ fdecl list \ 'c mdecl list \ 'a \ 'a) \ 'a" where "class_rec G == wfrec ((subcls1 G)^-1) (\r C t f. case class G C of None \ undefined diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/WellForm.thy --- a/src/HOL/MicroJava/J/WellForm.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/WellForm.thy Mon Mar 01 13:42:31 2010 +0100 @@ -27,45 +27,44 @@ *} types 'c wf_mb = "'c prog => cname => 'c mdecl => bool" -constdefs - wf_syscls :: "'c prog => bool" +definition wf_syscls :: "'c prog => bool" where "wf_syscls G == let cs = set G in Object \ fst ` cs \ (\x. Xcpt x \ fst ` cs)" - wf_fdecl :: "'c prog => fdecl => bool" +definition wf_fdecl :: "'c prog => fdecl => bool" where "wf_fdecl G == \(fn,ft). is_type G ft" - wf_mhead :: "'c prog => sig => ty => bool" +definition wf_mhead :: "'c prog => sig => ty => bool" where "wf_mhead G == \(mn,pTs) rT. (\T\set pTs. is_type G T) \ is_type G rT" - ws_cdecl :: "'c prog => 'c cdecl => bool" +definition ws_cdecl :: "'c prog => 'c cdecl => bool" where "ws_cdecl G == \(C,(D,fs,ms)). (\f\set fs. wf_fdecl G f) \ unique fs \ (\(sig,rT,mb)\set ms. wf_mhead G sig rT) \ unique ms \ (C \ Object \ is_class G D \ \G\D\C C)" - ws_prog :: "'c prog => bool" +definition ws_prog :: "'c prog => bool" where "ws_prog G == wf_syscls G \ (\c\set G. ws_cdecl G c) \ unique G" - wf_mrT :: "'c prog => 'c cdecl => bool" +definition wf_mrT :: "'c prog => 'c cdecl => bool" where "wf_mrT G == \(C,(D,fs,ms)). (C \ Object \ (\(sig,rT,b)\set ms. \D' rT' b'. method(G,D) sig = Some(D',rT',b') --> G\rT\rT'))" - wf_cdecl_mdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" +definition wf_cdecl_mdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" where "wf_cdecl_mdecl wf_mb G == \(C,(D,fs,ms)). (\m\set ms. wf_mb G C m)" - wf_prog :: "'c wf_mb => 'c prog => bool" +definition wf_prog :: "'c wf_mb => 'c prog => bool" where "wf_prog wf_mb G == ws_prog G \ (\c\ set G. wf_mrT G c \ wf_cdecl_mdecl wf_mb G c)" - wf_mdecl :: "'c wf_mb => 'c wf_mb" +definition wf_mdecl :: "'c wf_mb => 'c wf_mb" where "wf_mdecl wf_mb G C == \(sig,rT,mb). wf_mhead G sig rT \ wf_mb G C (sig,rT,mb)" - wf_cdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" +definition wf_cdecl :: "'c wf_mb => 'c prog => 'c cdecl => bool" where "wf_cdecl wf_mb G == \(C,(D,fs,ms)). (\f\set fs. wf_fdecl G f) \ unique fs \ diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/J/WellType.thy --- a/src/HOL/MicroJava/J/WellType.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/J/WellType.thy Mon Mar 01 13:42:31 2010 +0100 @@ -193,9 +193,7 @@ E\While(e) s\" -constdefs - - wf_java_mdecl :: "'c prog => cname => java_mb mdecl => bool" +definition wf_java_mdecl :: "'c prog => cname => java_mb mdecl => bool" where "wf_java_mdecl G C == \((mn,pTs),rT,(pns,lvars,blk,res)). length pTs = length pns \ distinct pns \ diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/JVM/JVMDefensive.thy --- a/src/HOL/MicroJava/JVM/JVMDefensive.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/JVM/JVMDefensive.thy Mon Mar 01 13:42:31 2010 +0100 @@ -4,7 +4,9 @@ header {* \isaheader{A Defensive JVM} *} -theory JVMDefensive imports JVMExec begin +theory JVMDefensive +imports JVMExec +begin text {* Extend the state space by one element indicating a type error (or @@ -16,39 +18,32 @@ fifth :: "'a \ 'b \ 'c \ 'd \ 'e \ 'f \ 'e" where "fifth x == fst(snd(snd(snd(snd x))))" - -consts isAddr :: "val \ bool" -recdef isAddr "{}" +fun isAddr :: "val \ bool" where "isAddr (Addr loc) = True" - "isAddr v = False" +| "isAddr v = False" -consts isIntg :: "val \ bool" -recdef isIntg "{}" +fun isIntg :: "val \ bool" where "isIntg (Intg i) = True" - "isIntg v = False" +| "isIntg v = False" -constdefs - isRef :: "val \ bool" +definition isRef :: "val \ bool" where "isRef v \ v = Null \ isAddr v" - -consts - check_instr :: "[instr, jvm_prog, aheap, opstack, locvars, - cname, sig, p_count, nat, frame list] \ bool" -primrec +primrec check_instr :: "[instr, jvm_prog, aheap, opstack, locvars, + cname, sig, p_count, nat, frame list] \ bool" where "check_instr (Load idx) G hp stk vars C sig pc mxs frs = (idx < length vars \ size stk < mxs)" - "check_instr (Store idx) G hp stk vars Cl sig pc mxs frs = +| "check_instr (Store idx) G hp stk vars Cl sig pc mxs frs = (0 < length stk \ idx < length vars)" - "check_instr (LitPush v) G hp stk vars Cl sig pc mxs frs = +| "check_instr (LitPush v) G hp stk vars Cl sig pc mxs frs = (\isAddr v \ size stk < mxs)" - "check_instr (New C) G hp stk vars Cl sig pc mxs frs = +| "check_instr (New C) G hp stk vars Cl sig pc mxs frs = (is_class G C \ size stk < mxs)" - "check_instr (Getfield F C) G hp stk vars Cl sig pc mxs frs = +| "check_instr (Getfield F C) G hp stk vars Cl sig pc mxs frs = (0 < length stk \ is_class G C \ field (G,C) F \ None \ (let (C', T) = the (field (G,C) F); ref = hd stk in C' = C \ isRef ref \ (ref \ Null \ @@ -56,7 +51,7 @@ (let (D,vs) = the (hp (the_Addr ref)) in G \ D \C C \ vs (F,C) \ None \ G,hp \ the (vs (F,C)) ::\ T))))" - "check_instr (Putfield F C) G hp stk vars Cl sig pc mxs frs = +| "check_instr (Putfield F C) G hp stk vars Cl sig pc mxs frs = (1 < length stk \ is_class G C \ field (G,C) F \ None \ (let (C', T) = the (field (G,C) F); v = hd stk; ref = hd (tl stk) in C' = C \ isRef ref \ (ref \ Null \ @@ -64,10 +59,10 @@ (let (D,vs) = the (hp (the_Addr ref)) in G \ D \C C \ G,hp \ v ::\ T))))" - "check_instr (Checkcast C) G hp stk vars Cl sig pc mxs frs = +| "check_instr (Checkcast C) G hp stk vars Cl sig pc mxs frs = (0 < length stk \ is_class G C \ isRef (hd stk))" - "check_instr (Invoke C mn ps) G hp stk vars Cl sig pc mxs frs = +| "check_instr (Invoke C mn ps) G hp stk vars Cl sig pc mxs frs = (length ps < length stk \ (let n = length ps; v = stk!n in isRef v \ (v \ Null \ @@ -75,41 +70,40 @@ method (G,cname_of hp v) (mn,ps) \ None \ list_all2 (\v T. G,hp \ v ::\ T) (rev (take n stk)) ps)))" - "check_instr Return G hp stk0 vars Cl sig0 pc mxs frs = +| "check_instr Return G hp stk0 vars Cl sig0 pc mxs frs = (0 < length stk0 \ (0 < length frs \ method (G,Cl) sig0 \ None \ (let v = hd stk0; (C, rT, body) = the (method (G,Cl) sig0) in Cl = C \ G,hp \ v ::\ rT)))" - "check_instr Pop G hp stk vars Cl sig pc mxs frs = +| "check_instr Pop G hp stk vars Cl sig pc mxs frs = (0 < length stk)" - "check_instr Dup G hp stk vars Cl sig pc mxs frs = +| "check_instr Dup G hp stk vars Cl sig pc mxs frs = (0 < length stk \ size stk < mxs)" - "check_instr Dup_x1 G hp stk vars Cl sig pc mxs frs = +| "check_instr Dup_x1 G hp stk vars Cl sig pc mxs frs = (1 < length stk \ size stk < mxs)" - "check_instr Dup_x2 G hp stk vars Cl sig pc mxs frs = +| "check_instr Dup_x2 G hp stk vars Cl sig pc mxs frs = (2 < length stk \ size stk < mxs)" - "check_instr Swap G hp stk vars Cl sig pc mxs frs = +| "check_instr Swap G hp stk vars Cl sig pc mxs frs = (1 < length stk)" - "check_instr IAdd G hp stk vars Cl sig pc mxs frs = +| "check_instr IAdd G hp stk vars Cl sig pc mxs frs = (1 < length stk \ isIntg (hd stk) \ isIntg (hd (tl stk)))" - "check_instr (Ifcmpeq b) G hp stk vars Cl sig pc mxs frs = +| "check_instr (Ifcmpeq b) G hp stk vars Cl sig pc mxs frs = (1 < length stk \ 0 \ int pc+b)" - "check_instr (Goto b) G hp stk vars Cl sig pc mxs frs = +| "check_instr (Goto b) G hp stk vars Cl sig pc mxs frs = (0 \ int pc+b)" - "check_instr Throw G hp stk vars Cl sig pc mxs frs = +| "check_instr Throw G hp stk vars Cl sig pc mxs frs = (0 < length stk \ isRef (hd stk))" -constdefs - check :: "jvm_prog \ jvm_state \ bool" +definition check :: "jvm_prog \ jvm_state \ bool" where "check G s \ let (xcpt, hp, frs) = s in (case frs of [] \ True | (stk,loc,C,sig,pc)#frs' \ (let (C',rt,mxs,mxl,ins,et) = the (method (G,C) sig); i = ins!pc in @@ -117,7 +111,7 @@ check_instr i G hp stk loc C sig pc mxs frs'))" - exec_d :: "jvm_prog \ jvm_state type_error \ jvm_state option type_error" +definition exec_d :: "jvm_prog \ jvm_state type_error \ jvm_state option type_error" where "exec_d G s \ case s of TypeError \ TypeError | Normal s' \ if check G s' then Normal (exec (G, s')) else TypeError" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/JVM/JVMExceptions.thy --- a/src/HOL/MicroJava/JVM/JVMExceptions.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/JVM/JVMExceptions.thy Mon Mar 01 13:42:31 2010 +0100 @@ -7,8 +7,7 @@ theory JVMExceptions imports JVMInstructions begin -constdefs - match_exception_entry :: "jvm_prog \ cname \ p_count \ exception_entry \ bool" +definition match_exception_entry :: "jvm_prog \ cname \ p_count \ exception_entry \ bool" where "match_exception_entry G cn pc ee == let (start_pc, end_pc, handler_pc, catch_type) = ee in start_pc <= pc \ pc < end_pc \ G\ cn \C catch_type" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/JVM/JVMExec.thy --- a/src/HOL/MicroJava/JVM/JVMExec.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/JVM/JVMExec.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/MicroJava/JVM/JVMExec.thy - ID: $Id$ Author: Cornelia Pusch, Gerwin Klein Copyright 1999 Technische Universitaet Muenchen *) @@ -26,9 +25,8 @@ "exec (G, Some xp, hp, frs) = None" -constdefs - exec_all :: "[jvm_prog,jvm_state,jvm_state] => bool" - ("_ |- _ -jvm-> _" [61,61,61]60) +definition exec_all :: "[jvm_prog,jvm_state,jvm_state] => bool" + ("_ |- _ -jvm-> _" [61,61,61]60) where "G |- s -jvm-> t == (s,t) \ {(s,t). exec(G,s) = Some t}^*" @@ -41,8 +39,7 @@ @{text this} pointer of the frame is set to @{text Null} to simulate a static method invokation. *} -constdefs - start_state :: "jvm_prog \ cname \ mname \ jvm_state" +definition start_state :: "jvm_prog \ cname \ mname \ jvm_state" where "start_state G C m \ let (C',rT,mxs,mxl,i,et) = the (method (G,C) (m,[])) in (None, start_heap G, [([], Null # replicate mxl undefined, C, (m,[]), 0)])" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/JVM/JVMListExample.thy --- a/src/HOL/MicroJava/JVM/JVMListExample.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/JVM/JVMListExample.thy Mon Mar 01 13:42:31 2010 +0100 @@ -16,20 +16,19 @@ val_nam :: vnam next_nam :: vnam -constdefs - list_name :: cname +definition list_name :: cname where "list_name == Cname list_nam" - test_name :: cname +definition test_name :: cname where "test_name == Cname test_nam" - val_name :: vname +definition val_name :: vname where "val_name == VName val_nam" - next_name :: vname +definition next_name :: vname where "next_name == VName next_nam" - append_ins :: bytecode +definition append_ins :: bytecode where "append_ins == [Load 0, Getfield next_name list_name, @@ -46,14 +45,14 @@ LitPush Unit, Return]" - list_class :: "jvm_method class" +definition list_class :: "jvm_method class" where "list_class == (Object, [(val_name, PrimT Integer), (next_name, Class list_name)], [((append_name, [Class list_name]), PrimT Void, (3, 0, append_ins,[(1,2,8,Xcpt NullPointer)]))])" - make_list_ins :: bytecode +definition make_list_ins :: bytecode where "make_list_ins == [New list_name, Dup, @@ -79,12 +78,12 @@ Invoke list_name append_name [Class list_name], Return]" - test_class :: "jvm_method class" +definition test_class :: "jvm_method class" where "test_class == (Object, [], [((makelist_name, []), PrimT Void, (3, 2, make_list_ins,[]))])" - E :: jvm_prog +definition E :: jvm_prog where "E == SystemClasses @ [(list_name, list_class), (test_name, test_class)]" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/MicroJava/JVM/JVMState.thy --- a/src/HOL/MicroJava/JVM/JVMState.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/MicroJava/JVM/JVMState.thy Mon Mar 01 13:42:31 2010 +0100 @@ -33,8 +33,7 @@ section {* Exceptions *} -constdefs - raise_system_xcpt :: "bool \ xcpt \ val option" +definition raise_system_xcpt :: "bool \ xcpt \ val option" where "raise_system_xcpt b x \ raise_if b x None" section {* Runtime State *} diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Modelcheck/CTL.thy --- a/src/HOL/Modelcheck/CTL.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Modelcheck/CTL.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Modelcheck/CTL.thy - ID: $Id$ Author: Olaf Mueller, Jan Philipps, Robert Sandner Copyright 1997 TU Muenchen *) @@ -11,10 +10,10 @@ types 'a trans = "('a * 'a) set" -constdefs - CEX ::"['a trans,'a pred, 'a]=>bool" +definition CEX :: "['a trans,'a pred, 'a]=>bool" where "CEX N f u == (? v. (f v & (u,v):N))" - EG ::"['a trans,'a pred]=> 'a pred" + +definition EG ::"['a trans,'a pred]=> 'a pred" where "EG N f == nu (% Q. % u.(f u & CEX N Q u))" end diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Modelcheck/EindhovenExample.thy --- a/src/HOL/Modelcheck/EindhovenExample.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Modelcheck/EindhovenExample.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Modelcheck/EindhovenExample.thy - ID: $Id$ Author: Olaf Mueller, Jan Philipps, Robert Sandner Copyright 1997 TU Muenchen *) @@ -11,17 +10,16 @@ types state = "bool * bool * bool" -constdefs - INIT :: "state pred" +definition INIT :: "state pred" where "INIT x == ~(fst x)&~(fst (snd x))&~(snd (snd x))" - N :: "[state,state] => bool" +definition N :: "[state,state] => bool" where "N == % (x1,x2,x3) (y1,y2,y3). (~x1 & ~x2 & ~x3 & y1 & ~y2 & ~y3) | ( x1 & ~x2 & ~x3 & ~y1 & ~y2 & ~y3) | ( x1 & ~x2 & ~x3 & y1 & y2 & y3)" - reach:: "state pred" +definition reach:: "state pred" where "reach == mu (%Q x. INIT x | (? y. Q y & N y x))" lemma init_state: "INIT (a, b, c) = (~a & ~b &~c)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Modelcheck/MuCalculus.thy --- a/src/HOL/Modelcheck/MuCalculus.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Modelcheck/MuCalculus.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Modelcheck/MuCalculus.thy - ID: $Id$ Author: Olaf Mueller, Jan Philipps, Robert Sandner Copyright 1997 TU Muenchen *) @@ -11,17 +10,16 @@ types 'a pred = "'a=>bool" -constdefs - Charfun :: "'a set => 'a pred" +definition Charfun :: "'a set => 'a pred" where "Charfun == (% A.% x. x:A)" - monoP :: "('a pred => 'a pred) => bool" +definition monoP :: "('a pred => 'a pred) => bool" where "monoP f == mono(Collect o f o Charfun)" - mu :: "('a pred => 'a pred) => 'a pred" (binder "Mu " 10) +definition mu :: "('a pred => 'a pred) => 'a pred" (binder "Mu " 10) where "mu f == Charfun(lfp(Collect o f o Charfun))" - nu :: "('a pred => 'a pred) => 'a pred" (binder "Nu " 10) +definition nu :: "('a pred => 'a pred) => 'a pred" (binder "Nu " 10) where "nu f == Charfun(gfp(Collect o f o Charfun))" end diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Modelcheck/MuckeExample1.thy --- a/src/HOL/Modelcheck/MuckeExample1.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Modelcheck/MuckeExample1.thy Mon Mar 01 13:42:31 2010 +0100 @@ -11,8 +11,7 @@ types state = "bool * bool * bool" -constdefs - INIT :: "state pred" +definition INIT :: "state pred" where "INIT x == ~(fst x)&~(fst (snd x))&~(snd (snd x))" N :: "[state,state] => bool" "N x y == let x1 = fst(x); x2 = fst(snd(x)); x3 = snd(snd(x)); diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Modelcheck/MuckeExample2.thy --- a/src/HOL/Modelcheck/MuckeExample2.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Modelcheck/MuckeExample2.thy Mon Mar 01 13:42:31 2010 +0100 @@ -8,8 +8,7 @@ imports MuckeSyn begin -constdefs - Init :: "bool pred" +definition Init :: "bool pred" where "Init x == x" R :: "[bool,bool] => bool" "R x y == (x & ~y) | (~x & y)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/NanoJava/Decl.thy --- a/src/HOL/NanoJava/Decl.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/NanoJava/Decl.thy Mon Mar 01 13:42:31 2010 +0100 @@ -50,11 +50,10 @@ Object :: cname --{* name of root class *} -constdefs - "class" :: "cname \ class" +definition "class" :: "cname \ class" where "class \ map_of Prog" - is_class :: "cname => bool" +definition is_class :: "cname => bool" where "is_class C \ class C \ None" lemma finite_is_class: "finite {C. is_class C}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/NanoJava/Equivalence.thy --- a/src/HOL/NanoJava/Equivalence.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/NanoJava/Equivalence.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/NanoJava/Equivalence.thy - ID: $Id$ Author: David von Oheimb Copyright 2001 Technische Universitaet Muenchen *) @@ -10,27 +9,25 @@ subsection "Validity" -constdefs - valid :: "[assn,stmt, assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) +definition valid :: "[assn,stmt, assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) where "|= {P} c {Q} \ \s t. P s --> (\n. s -c -n\ t) --> Q t" - evalid :: "[assn,expr,vassn] => bool" ("|=e {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) +definition evalid :: "[assn,expr,vassn] => bool" ("|=e {(1_)}/ (_)/ {(1_)}" [3,90,3] 60) where "|=e {P} e {Q} \ \s v t. P s --> (\n. s -e\v-n\ t) --> Q v t" - - nvalid :: "[nat, triple ] => bool" ("|=_: _" [61,61] 60) +definition nvalid :: "[nat, triple ] => bool" ("|=_: _" [61,61] 60) where "|=n: t \ let (P,c,Q) = t in \s t. s -c -n\ t --> P s --> Q t" -envalid :: "[nat,etriple ] => bool" ("|=_:e _" [61,61] 60) +definition envalid :: "[nat,etriple ] => bool" ("|=_:e _" [61,61] 60) where "|=n:e t \ let (P,e,Q) = t in \s v t. s -e\v-n\ t --> P s --> Q v t" - nvalids :: "[nat, triple set] => bool" ("||=_: _" [61,61] 60) +definition nvalids :: "[nat, triple set] => bool" ("||=_: _" [61,61] 60) where "||=n: T \ \t\T. |=n: t" - cnvalids :: "[triple set,triple set] => bool" ("_ ||=/ _" [61,61] 60) +definition cnvalids :: "[triple set,triple set] => bool" ("_ ||=/ _" [61,61] 60) where "A ||= C \ \n. ||=n: A --> ||=n: C" -cenvalid :: "[triple set,etriple ] => bool" ("_ ||=e/ _" [61,61] 60) +definition cenvalid :: "[triple set,etriple ] => bool" ("_ ||=e/ _" [61,61] 60) where "A ||=e t \ \n. ||=n: A --> |=n:e t" notation (xsymbols) @@ -160,10 +157,12 @@ subsection "(Relative) Completeness" -constdefs MGT :: "stmt => state => triple" +definition MGT :: "stmt => state => triple" where "MGT c Z \ (\s. Z = s, c, \ t. \n. Z -c- n\ t)" - MGTe :: "expr => state => etriple" + +definition MGTe :: "expr => state => etriple" where "MGTe e Z \ (\s. Z = s, e, \v t. \n. Z -e\v-n\ t)" + notation (xsymbols) MGTe ("MGT\<^sub>e") notation (HTML output) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/NanoJava/State.thy --- a/src/HOL/NanoJava/State.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/NanoJava/State.thy Mon Mar 01 13:42:31 2010 +0100 @@ -7,9 +7,7 @@ theory State imports TypeRel begin -constdefs - - body :: "cname \ mname => stmt" +definition body :: "cname \ mname => stmt" where "body \ \(C,m). bdy (the (method C m))" text {* Locations, i.e.\ abstract references to objects *} @@ -29,9 +27,7 @@ "fields" \ (type)"fname => val option" "obj" \ (type)"cname \ fields" -constdefs - - init_vars:: "('a \ 'b) => ('a \ val)" +definition init_vars :: "('a \ 'b) => ('a \ val)" where "init_vars m == Option.map (\T. Null) o m" text {* private: *} @@ -49,54 +45,49 @@ "locals" \ (type)"vname => val option" "state" \ (type)"(|heap :: heap, locals :: locals|)" -constdefs - - del_locs :: "state => state" +definition del_locs :: "state => state" where "del_locs s \ s (| locals := empty |)" - init_locs :: "cname => mname => state => state" +definition init_locs :: "cname => mname => state => state" where "init_locs C m s \ s (| locals := locals s ++ init_vars (map_of (lcl (the (method C m)))) |)" text {* The first parameter of @{term set_locs} is of type @{typ state} rather than @{typ locals} in order to keep @{typ locals} private.*} -constdefs - set_locs :: "state => state => state" +definition set_locs :: "state => state => state" where "set_locs s s' \ s' (| locals := locals s |)" - get_local :: "state => vname => val" ("_<_>" [99,0] 99) +definition get_local :: "state => vname => val" ("_<_>" [99,0] 99) where "get_local s x \ the (locals s x)" --{* local function: *} - get_obj :: "state => loc => obj" +definition get_obj :: "state => loc => obj" where "get_obj s a \ the (heap s a)" - obj_class :: "state => loc => cname" +definition obj_class :: "state => loc => cname" where "obj_class s a \ fst (get_obj s a)" - get_field :: "state => loc => fname => val" +definition get_field :: "state => loc => fname => val" where "get_field s a f \ the (snd (get_obj s a) f)" --{* local function: *} - hupd :: "loc => obj => state => state" ("hupd'(_|->_')" [10,10] 1000) +definition hupd :: "loc => obj => state => state" ("hupd'(_|->_')" [10,10] 1000) where "hupd a obj s \ s (| heap := ((heap s)(a\obj))|)" - lupd :: "vname => val => state => state" ("lupd'(_|->_')" [10,10] 1000) +definition lupd :: "vname => val => state => state" ("lupd'(_|->_')" [10,10] 1000) where "lupd x v s \ s (| locals := ((locals s)(x\v ))|)" notation (xsymbols) hupd ("hupd'(_\_')" [10,10] 1000) and lupd ("lupd'(_\_')" [10,10] 1000) -constdefs - - new_obj :: "loc => cname => state => state" +definition new_obj :: "loc => cname => state => state" where "new_obj a C \ hupd(a\(C,init_vars (field C)))" - upd_obj :: "loc => fname => val => state => state" +definition upd_obj :: "loc => fname => val => state => state" where "upd_obj a f v s \ let (C,fs) = the (heap s a) in hupd(a\(C,fs(f\v))) s" - new_Addr :: "state => val" +definition new_Addr :: "state => val" where "new_Addr s == SOME v. (\a. v = Addr a \ (heap s) a = None) | v = Null" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/NanoJava/TypeRel.thy --- a/src/HOL/NanoJava/TypeRel.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/NanoJava/TypeRel.thy Mon Mar 01 13:42:31 2010 +0100 @@ -66,9 +66,7 @@ apply auto done -constdefs - - ws_prog :: "bool" +definition ws_prog :: "bool" where "ws_prog \ \(C,c)\set Prog. C\Object \ is_class (super c) \ (super c,C)\subcls1^+" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Nat.thy --- a/src/HOL/Nat.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Nat.thy Mon Mar 01 13:42:31 2010 +0100 @@ -45,8 +45,7 @@ nat = Nat by (rule exI, unfold mem_def, rule Nat.Zero_RepI) -constdefs - Suc :: "nat => nat" +definition Suc :: "nat => nat" where Suc_def: "Suc == (%n. Abs_Nat (Suc_Rep (Rep_Nat n)))" local diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Nitpick_Examples/Refute_Nits.thy --- a/src/HOL/Nitpick_Examples/Refute_Nits.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Nitpick_Examples/Refute_Nits.thy Mon Mar 01 13:42:31 2010 +0100 @@ -244,8 +244,7 @@ text {* ``The transitive closure of an arbitrary relation is non-empty.'' *} -constdefs -"trans" :: "('a \ 'a \ bool) \ bool" +definition "trans" :: "('a \ 'a \ bool) \ bool" where "trans P \ (ALL x y z. P x y \ P y z \ P x z)" "subset" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" "subset P Q \ (ALL x y. P x y \ Q x y)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Nominal/Examples/Class.thy --- a/src/HOL/Nominal/Examples/Class.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Nominal/Examples/Class.thy Mon Mar 01 13:42:31 2010 +0100 @@ -10289,17 +10289,16 @@ text {* set operators *} -constdefs - AXIOMSn::"ty \ ntrm set" +definition AXIOMSn :: "ty \ ntrm set" where "AXIOMSn B \ { (x):(Ax y b) | x y b. True }" - AXIOMSc::"ty \ ctrm set" +definition AXIOMSc::"ty \ ctrm set" where "AXIOMSc B \ { :(Ax y b) | a y b. True }" - BINDINGn::"ty \ ctrm set \ ntrm set" +definition BINDINGn::"ty \ ctrm set \ ntrm set" where "BINDINGn B X \ { (x):M | x M. \a P. :P\X \ SNa (M{x:=.P})}" - BINDINGc::"ty \ ntrm set \ ctrm set" +definition BINDINGc::"ty \ ntrm set \ ctrm set" where "BINDINGc B X \ { :M | a M. \x P. (x):P\X \ SNa (M{a:=(x).P})}" lemma BINDINGn_decreasing: @@ -16540,8 +16539,7 @@ apply(fast) done -constdefs - SNa_Redu :: "(trm \ trm) set" +definition SNa_Redu :: "(trm \ trm) set" where "SNa_Redu \ A_Redu_set \ (UNIV <*> SNa_set)" lemma wf_SNa_Redu: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Nominal/Examples/Fsub.thy --- a/src/HOL/Nominal/Examples/Fsub.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Nominal/Examples/Fsub.thy Mon Mar 01 13:42:31 2010 +0100 @@ -223,8 +223,7 @@ in @{term "\"}. The set of free variables of @{term "S"} is the @{text "support"} of @{term "S"}. *} -constdefs - "closed_in" :: "ty \ env \ bool" ("_ closed'_in _" [100,100] 100) +definition "closed_in" :: "ty \ env \ bool" ("_ closed'_in _" [100,100] 100) where "S closed_in \ \ (supp S)\(ty_dom \)" lemma closed_in_eqvt[eqvt]: @@ -687,13 +686,13 @@ have fresh_cond: "X\\" by fact hence fresh_ty_dom: "X\(ty_dom \)" by (simp add: fresh_dom) have "(\X<:T\<^isub>2. T\<^isub>1) closed_in \" by fact - hence closed\<^isub>T\<^isub>2: "T\<^isub>2 closed_in \" and closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in ((TVarB X T\<^isub>2)#\)" + hence closed\<^isub>T2: "T\<^isub>2 closed_in \" and closed\<^isub>T1: "T\<^isub>1 closed_in ((TVarB X T\<^isub>2)#\)" by (auto simp add: closed_in_def ty.supp abs_supp) have ok: "\ \ ok" by fact - hence ok': "\ ((TVarB X T\<^isub>2)#\) ok" using closed\<^isub>T\<^isub>2 fresh_ty_dom by simp - have "\ \ T\<^isub>2 <: T\<^isub>2" using ih_T\<^isub>2 closed\<^isub>T\<^isub>2 ok by simp + hence ok': "\ ((TVarB X T\<^isub>2)#\) ok" using closed\<^isub>T2 fresh_ty_dom by simp + have "\ \ T\<^isub>2 <: T\<^isub>2" using ih_T\<^isub>2 closed\<^isub>T2 ok by simp moreover - have "((TVarB X T\<^isub>2)#\) \ T\<^isub>1 <: T\<^isub>1" using ih_T\<^isub>1 closed\<^isub>T\<^isub>1 ok' by simp + have "((TVarB X T\<^isub>2)#\) \ T\<^isub>1 <: T\<^isub>1" using ih_T\<^isub>1 closed\<^isub>T1 ok' by simp ultimately show "\ \ (\X<:T\<^isub>2. T\<^isub>1) <: (\X<:T\<^isub>2. T\<^isub>1)" using fresh_cond by (simp add: subtype_of.SA_all) qed (auto simp add: closed_in_def ty.supp supp_atm) @@ -718,8 +717,7 @@ another. This generalization seems to make the proof for weakening to be smoother than if we had strictly adhered to the version in the POPLmark-paper. *} -constdefs - extends :: "env \ env \ bool" ("_ extends _" [100,100] 100) +definition extends :: "env \ env \ bool" ("_ extends _" [100,100] 100) where "\ extends \ \ \X Q. (TVarB X Q)\set \ \ (TVarB X Q)\set \" lemma extends_ty_dom: @@ -785,10 +783,10 @@ have ih\<^isub>1: "\\. \ \ ok \ \ extends \ \ \ \ T\<^isub>1 <: S\<^isub>1" by fact have ih\<^isub>2: "\\. \ \ ok \ \ extends ((TVarB X T\<^isub>1)#\) \ \ \ S\<^isub>2 <: T\<^isub>2" by fact have lh_drv_prem: "\ \ T\<^isub>1 <: S\<^isub>1" by fact - hence closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in \" by (simp add: subtype_implies_closed) + hence closed\<^isub>T1: "T\<^isub>1 closed_in \" by (simp add: subtype_implies_closed) have ok: "\ \ ok" by fact have ext: "\ extends \" by fact - have "T\<^isub>1 closed_in \" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) + have "T\<^isub>1 closed_in \" using ext closed\<^isub>T1 by (simp only: extends_closed) hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_dom ok by force moreover have "((TVarB X T\<^isub>1)#\) extends ((TVarB X T\<^isub>1)#\)" using ext by (force simp add: extends_def) @@ -813,10 +811,10 @@ have ih\<^isub>1: "\\. \ \ ok \ \ extends \ \ \ \ T\<^isub>1 <: S\<^isub>1" by fact have ih\<^isub>2: "\\. \ \ ok \ \ extends ((TVarB X T\<^isub>1)#\) \ \ \ S\<^isub>2 <: T\<^isub>2" by fact have lh_drv_prem: "\ \ T\<^isub>1 <: S\<^isub>1" by fact - hence closed\<^isub>T\<^isub>1: "T\<^isub>1 closed_in \" by (simp add: subtype_implies_closed) + hence closed\<^isub>T1: "T\<^isub>1 closed_in \" by (simp add: subtype_implies_closed) have ok: "\ \ ok" by fact have ext: "\ extends \" by fact - have "T\<^isub>1 closed_in \" using ext closed\<^isub>T\<^isub>1 by (simp only: extends_closed) + have "T\<^isub>1 closed_in \" using ext closed\<^isub>T1 by (simp only: extends_closed) hence "\ ((TVarB X T\<^isub>1)#\) ok" using fresh_dom ok by force moreover have "((TVarB X T\<^isub>1)#\) extends ((TVarB X T\<^isub>1)#\)" using ext by (force simp add: extends_def) @@ -905,7 +903,7 @@ case (SA_arrow \ Q\<^isub>1 S\<^isub>1 S\<^isub>2 Q\<^isub>2) then have rh_drv: "\ \ Q\<^isub>1 \ Q\<^isub>2 <: T" by simp from `Q\<^isub>1 \ Q\<^isub>2 = Q` - have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" by auto + have Q\<^isub>12_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" by auto have lh_drv_prm\<^isub>1: "\ \ Q\<^isub>1 <: S\<^isub>1" by fact have lh_drv_prm\<^isub>2: "\ \ S\<^isub>2 <: Q\<^isub>2" by fact from rh_drv have "T=Top \ (\T\<^isub>1 T\<^isub>2. T=T\<^isub>1\T\<^isub>2 \ \\T\<^isub>1<:Q\<^isub>1 \ \\Q\<^isub>2<:T\<^isub>2)" @@ -923,10 +921,10 @@ and rh_drv_prm\<^isub>1: "\ \ T\<^isub>1 <: Q\<^isub>1" and rh_drv_prm\<^isub>2: "\ \ Q\<^isub>2 <: T\<^isub>2" by force from IH_trans[of "Q\<^isub>1"] - have "\ \ T\<^isub>1 <: S\<^isub>1" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>1 lh_drv_prm\<^isub>1 by simp + have "\ \ T\<^isub>1 <: S\<^isub>1" using Q\<^isub>12_less rh_drv_prm\<^isub>1 lh_drv_prm\<^isub>1 by simp moreover from IH_trans[of "Q\<^isub>2"] - have "\ \ S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 lh_drv_prm\<^isub>2 by simp + have "\ \ S\<^isub>2 <: T\<^isub>2" using Q\<^isub>12_less rh_drv_prm\<^isub>2 lh_drv_prm\<^isub>2 by simp ultimately have "\ \ S\<^isub>1 \ S\<^isub>2 <: T\<^isub>1 \ T\<^isub>2" by auto then have "\ \ S\<^isub>1 \ S\<^isub>2 <: T" using T_inst by simp } @@ -956,15 +954,15 @@ and rh_drv_prm\<^isub>1: "\ \ T\<^isub>1 <: Q\<^isub>1" and rh_drv_prm\<^isub>2:"((TVarB X T\<^isub>1)#\) \ Q\<^isub>2 <: T\<^isub>2" by force have "(\X<:Q\<^isub>1. Q\<^isub>2) = Q" by fact - then have Q\<^isub>1\<^isub>2_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" + then have Q\<^isub>12_less: "size_ty Q\<^isub>1 < size_ty Q" "size_ty Q\<^isub>2 < size_ty Q" using fresh_cond by auto from IH_trans[of "Q\<^isub>1"] - have "\ \ T\<^isub>1 <: S\<^isub>1" using lh_drv_prm\<^isub>1 rh_drv_prm\<^isub>1 Q\<^isub>1\<^isub>2_less by blast + have "\ \ T\<^isub>1 <: S\<^isub>1" using lh_drv_prm\<^isub>1 rh_drv_prm\<^isub>1 Q\<^isub>12_less by blast moreover from IH_narrow[of "Q\<^isub>1" "[]"] - have "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: Q\<^isub>2" using Q\<^isub>1\<^isub>2_less lh_drv_prm\<^isub>2 rh_drv_prm\<^isub>1 by simp + have "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: Q\<^isub>2" using Q\<^isub>12_less lh_drv_prm\<^isub>2 rh_drv_prm\<^isub>1 by simp with IH_trans[of "Q\<^isub>2"] - have "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2" using Q\<^isub>1\<^isub>2_less rh_drv_prm\<^isub>2 by simp + have "((TVarB X T\<^isub>1)#\) \ S\<^isub>2 <: T\<^isub>2" using Q\<^isub>12_less rh_drv_prm\<^isub>2 by simp ultimately have "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: (\X<:T\<^isub>1. T\<^isub>2)" using fresh_cond by (simp add: subtype_of.SA_all) hence "\ \ (\X<:S\<^isub>1. S\<^isub>2) <: T" using T_inst by simp @@ -1007,16 +1005,16 @@ with IH_inner show "(\@[(TVarB X P)]@\) \ Tvar Y <: N" by (simp add: subtype_of.SA_trans_TVar) next case True - have memb\<^isub>X\<^isub>Q: "(TVarB X Q)\set (\@[(TVarB X Q)]@\)" by simp - have memb\<^isub>X\<^isub>P: "(TVarB X P)\set (\@[(TVarB X P)]@\)" by simp + have memb\<^isub>XQ: "(TVarB X Q)\set (\@[(TVarB X Q)]@\)" by simp + have memb\<^isub>XP: "(TVarB X P)\set (\@[(TVarB X P)]@\)" by simp have eq: "X=Y" by fact - hence "S=Q" using ok\<^isub>Q lh_drv_prm memb\<^isub>X\<^isub>Q by (simp only: uniqueness_of_ctxt) + hence "S=Q" using ok\<^isub>Q lh_drv_prm memb\<^isub>XQ by (simp only: uniqueness_of_ctxt) hence "(\@[(TVarB X P)]@\) \ Q <: N" using IH_inner by simp moreover have "(\@[(TVarB X P)]@\) extends \" by (simp add: extends_def) hence "(\@[(TVarB X P)]@\) \ P <: Q" using rh_drv ok\<^isub>P by (simp only: weakening) ultimately have "(\@[(TVarB X P)]@\) \ P <: N" by (simp add: transitivity_lemma) - then show "(\@[(TVarB X P)]@\) \ Tvar Y <: N" using memb\<^isub>X\<^isub>P eq by auto + then show "(\@[(TVarB X P)]@\) \ Tvar Y <: N" using memb\<^isub>XP eq by auto qed next case (SA_refl_TVar Y \ X \) @@ -1051,7 +1049,7 @@ | T_Abs[intro]: "\ VarB x T\<^isub>1 # \ \ t\<^isub>2 : T\<^isub>2 \ \ \ \ (\x:T\<^isub>1. t\<^isub>2) : T\<^isub>1 \ T\<^isub>2" | T_Sub[intro]: "\ \ \ t : S; \ \ S <: T \ \ \ \ t : T" | T_TAbs[intro]:"\ TVarB X T\<^isub>1 # \ \ t\<^isub>2 : T\<^isub>2 \ \ \ \ (\X<:T\<^isub>1. t\<^isub>2) : (\X<:T\<^isub>1. T\<^isub>2)" -| T_TApp[intro]:"\X\(\,t\<^isub>1,T\<^isub>2); \ \ t\<^isub>1 : (\X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2); \ \ T\<^isub>2 <: T\<^isub>1\<^isub>1\ \ \ \ t\<^isub>1 \\<^sub>\ T\<^isub>2 : (T\<^isub>1\<^isub>2[X \ T\<^isub>2]\<^sub>\)" +| T_TApp[intro]:"\X\(\,t\<^isub>1,T\<^isub>2); \ \ t\<^isub>1 : (\X<:T\<^isub>11. T\<^isub>12); \ \ T\<^isub>2 <: T\<^isub>11\ \ \ \ t\<^isub>1 \\<^sub>\ T\<^isub>2 : (T\<^isub>12[X \ T\<^isub>2]\<^sub>\)" equivariance typing @@ -1166,10 +1164,10 @@ inductive eval :: "trm \ trm \ bool" ("_ \ _" [60,60] 60) where - E_Abs : "\ x \ v\<^isub>2; val v\<^isub>2 \ \ (\x:T\<^isub>1\<^isub>1. t\<^isub>1\<^isub>2) \ v\<^isub>2 \ t\<^isub>1\<^isub>2[x \ v\<^isub>2]" + E_Abs : "\ x \ v\<^isub>2; val v\<^isub>2 \ \ (\x:T\<^isub>11. t\<^isub>12) \ v\<^isub>2 \ t\<^isub>12[x \ v\<^isub>2]" | E_App1 [intro]: "t \ t' \ t \ u \ t' \ u" | E_App2 [intro]: "\ val v; t \ t' \ \ v \ t \ v \ t'" -| E_TAbs : "X \ (T\<^isub>1\<^isub>1, T\<^isub>2) \ (\X<:T\<^isub>1\<^isub>1. t\<^isub>1\<^isub>2) \\<^sub>\ T\<^isub>2 \ t\<^isub>1\<^isub>2[X \\<^sub>\ T\<^isub>2]" +| E_TAbs : "X \ (T\<^isub>11, T\<^isub>2) \ (\X<:T\<^isub>11. t\<^isub>12) \\<^sub>\ T\<^isub>2 \ t\<^isub>12[X \\<^sub>\ T\<^isub>2]" | E_TApp [intro]: "t \ t' \ t \\<^sub>\ T \ t' \\<^sub>\ T" lemma better_E_Abs[intro]: @@ -1317,7 +1315,7 @@ case (T_Var x T) then show ?case by auto next - case (T_App X t\<^isub>1 T\<^isub>2 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2) + case (T_App X t\<^isub>1 T\<^isub>2 T\<^isub>11 T\<^isub>12) then show ?case by force next case (T_Abs y T\<^isub>1 t\<^isub>2 T\<^isub>2 \ \) @@ -1746,68 +1744,68 @@ assumes H: "\ \ t : T" shows "t \ t' \ \ \ t' : T" using H proof (nominal_induct avoiding: t' rule: typing.strong_induct) - case (T_App \ t\<^isub>1 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2 t\<^isub>2 t') + case (T_App \ t\<^isub>1 T\<^isub>11 T\<^isub>12 t\<^isub>2 t') obtain x::vrs where x_fresh: "x \ (\, t\<^isub>1 \ t\<^isub>2, t')" by (rule exists_fresh) (rule fin_supp) obtain X::tyvrs where "X \ (t\<^isub>1 \ t\<^isub>2, t')" by (rule exists_fresh) (rule fin_supp) with `t\<^isub>1 \ t\<^isub>2 \ t'` show ?case proof (cases rule: eval.strong_cases [where x=x and X=X]) - case (E_Abs v\<^isub>2 T\<^isub>1\<^isub>1' t\<^isub>1\<^isub>2) - with T_App and x_fresh have h: "\ \ (\x:T\<^isub>1\<^isub>1'. t\<^isub>1\<^isub>2) : T\<^isub>1\<^isub>1 \ T\<^isub>1\<^isub>2" + case (E_Abs v\<^isub>2 T\<^isub>11' t\<^isub>12) + with T_App and x_fresh have h: "\ \ (\x:T\<^isub>11'. t\<^isub>12) : T\<^isub>11 \ T\<^isub>12" by (simp add: trm.inject fresh_prod) moreover from x_fresh have "x \ \" by simp ultimately obtain S' - where T\<^isub>1\<^isub>1: "\ \ T\<^isub>1\<^isub>1 <: T\<^isub>1\<^isub>1'" - and t\<^isub>1\<^isub>2: "(VarB x T\<^isub>1\<^isub>1') # \ \ t\<^isub>1\<^isub>2 : S'" - and S': "\ \ S' <: T\<^isub>1\<^isub>2" + where T\<^isub>11: "\ \ T\<^isub>11 <: T\<^isub>11'" + and t\<^isub>12: "(VarB x T\<^isub>11') # \ \ t\<^isub>12 : S'" + and S': "\ \ S' <: T\<^isub>12" by (rule Abs_type') blast - from `\ \ t\<^isub>2 : T\<^isub>1\<^isub>1` - have "\ \ t\<^isub>2 : T\<^isub>1\<^isub>1'" using T\<^isub>1\<^isub>1 by (rule T_Sub) - with t\<^isub>1\<^isub>2 have "\ \ t\<^isub>1\<^isub>2[x \ t\<^isub>2] : S'" + from `\ \ t\<^isub>2 : T\<^isub>11` + have "\ \ t\<^isub>2 : T\<^isub>11'" using T\<^isub>11 by (rule T_Sub) + with t\<^isub>12 have "\ \ t\<^isub>12[x \ t\<^isub>2] : S'" by (rule subst_type [where \="[]", simplified]) - hence "\ \ t\<^isub>1\<^isub>2[x \ t\<^isub>2] : T\<^isub>1\<^isub>2" using S' by (rule T_Sub) + hence "\ \ t\<^isub>12[x \ t\<^isub>2] : T\<^isub>12" using S' by (rule T_Sub) with E_Abs and x_fresh show ?thesis by (simp add: trm.inject fresh_prod) next case (E_App1 t''' t'' u) hence "t\<^isub>1 \ t''" by (simp add:trm.inject) - hence "\ \ t'' : T\<^isub>1\<^isub>1 \ T\<^isub>1\<^isub>2" by (rule T_App) - hence "\ \ t'' \ t\<^isub>2 : T\<^isub>1\<^isub>2" using `\ \ t\<^isub>2 : T\<^isub>1\<^isub>1` + hence "\ \ t'' : T\<^isub>11 \ T\<^isub>12" by (rule T_App) + hence "\ \ t'' \ t\<^isub>2 : T\<^isub>12" using `\ \ t\<^isub>2 : T\<^isub>11` by (rule typing.T_App) with E_App1 show ?thesis by (simp add:trm.inject) next case (E_App2 v t''' t'') hence "t\<^isub>2 \ t''" by (simp add:trm.inject) - hence "\ \ t'' : T\<^isub>1\<^isub>1" by (rule T_App) - with T_App(1) have "\ \ t\<^isub>1 \ t'' : T\<^isub>1\<^isub>2" + hence "\ \ t'' : T\<^isub>11" by (rule T_App) + with T_App(1) have "\ \ t\<^isub>1 \ t'' : T\<^isub>12" by (rule typing.T_App) with E_App2 show ?thesis by (simp add:trm.inject) qed (simp_all add: fresh_prod) next - case (T_TApp X \ t\<^isub>1 T\<^isub>2 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2 t') + case (T_TApp X \ t\<^isub>1 T\<^isub>2 T\<^isub>11 T\<^isub>12 t') obtain x::vrs where "x \ (t\<^isub>1 \\<^sub>\ T\<^isub>2, t')" by (rule exists_fresh) (rule fin_supp) with `t\<^isub>1 \\<^sub>\ T\<^isub>2 \ t'` show ?case proof (cases rule: eval.strong_cases [where X=X and x=x]) - case (E_TAbs T\<^isub>1\<^isub>1' T\<^isub>2' t\<^isub>1\<^isub>2) - with T_TApp have "\ \ (\X<:T\<^isub>1\<^isub>1'. t\<^isub>1\<^isub>2) : (\X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2)" and "X \ \" and "X \ T\<^isub>1\<^isub>1'" + case (E_TAbs T\<^isub>11' T\<^isub>2' t\<^isub>12) + with T_TApp have "\ \ (\X<:T\<^isub>11'. t\<^isub>12) : (\X<:T\<^isub>11. T\<^isub>12)" and "X \ \" and "X \ T\<^isub>11'" by (simp_all add: trm.inject) - moreover from `\\T\<^isub>2<:T\<^isub>1\<^isub>1` and `X \ \` have "X \ T\<^isub>1\<^isub>1" + moreover from `\\T\<^isub>2<:T\<^isub>11` and `X \ \` have "X \ T\<^isub>11" by (blast intro: closed_in_fresh fresh_dom dest: subtype_implies_closed) ultimately obtain S' - where "TVarB X T\<^isub>1\<^isub>1 # \ \ t\<^isub>1\<^isub>2 : S'" - and "(TVarB X T\<^isub>1\<^isub>1 # \) \ S' <: T\<^isub>1\<^isub>2" + where "TVarB X T\<^isub>11 # \ \ t\<^isub>12 : S'" + and "(TVarB X T\<^isub>11 # \) \ S' <: T\<^isub>12" by (rule TAbs_type') blast - hence "TVarB X T\<^isub>1\<^isub>1 # \ \ t\<^isub>1\<^isub>2 : T\<^isub>1\<^isub>2" by (rule T_Sub) - hence "\ \ t\<^isub>1\<^isub>2[X \\<^sub>\ T\<^isub>2] : T\<^isub>1\<^isub>2[X \ T\<^isub>2]\<^sub>\" using `\ \ T\<^isub>2 <: T\<^isub>1\<^isub>1` + hence "TVarB X T\<^isub>11 # \ \ t\<^isub>12 : T\<^isub>12" by (rule T_Sub) + hence "\ \ t\<^isub>12[X \\<^sub>\ T\<^isub>2] : T\<^isub>12[X \ T\<^isub>2]\<^sub>\" using `\ \ T\<^isub>2 <: T\<^isub>11` by (rule substT_type [where D="[]", simplified]) with T_TApp and E_TAbs show ?thesis by (simp add: trm.inject) next case (E_TApp t''' t'' T) from E_TApp have "t\<^isub>1 \ t''" by (simp add: trm.inject) - then have "\ \ t'' : (\X<:T\<^isub>1\<^isub>1. T\<^isub>1\<^isub>2)" by (rule T_TApp) - then have "\ \ t'' \\<^sub>\ T\<^isub>2 : T\<^isub>1\<^isub>2[X \ T\<^isub>2]\<^sub>\" using `\ \ T\<^isub>2 <: T\<^isub>1\<^isub>1` + then have "\ \ t'' : (\X<:T\<^isub>11. T\<^isub>12)" by (rule T_TApp) + then have "\ \ t'' \\<^sub>\ T\<^isub>2 : T\<^isub>12[X \ T\<^isub>2]\<^sub>\" using `\ \ T\<^isub>2 <: T\<^isub>11` by (rule better_T_TApp) with E_TApp show ?thesis by (simp add: trm.inject) qed (simp_all add: fresh_prod) @@ -1847,7 +1845,7 @@ shows "val t \ (\t'. t \ t')" using assms proof (induct "[]::env" t T) - case (T_App t\<^isub>1 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2 t\<^isub>2) + case (T_App t\<^isub>1 T\<^isub>11 T\<^isub>12 t\<^isub>2) hence "val t\<^isub>1 \ (\t'. t\<^isub>1 \ t')" by simp thus ?case proof @@ -1873,7 +1871,7 @@ thus ?case by auto qed next - case (T_TApp X t\<^isub>1 T\<^isub>2 T\<^isub>1\<^isub>1 T\<^isub>1\<^isub>2) + case (T_TApp X t\<^isub>1 T\<^isub>2 T\<^isub>11 T\<^isub>12) hence "val t\<^isub>1 \ (\t'. t\<^isub>1 \ t')" by simp thus ?case proof diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Nominal/Examples/LocalWeakening.thy --- a/src/HOL/Nominal/Examples/LocalWeakening.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Nominal/Examples/LocalWeakening.thy Mon Mar 01 13:42:31 2010 +0100 @@ -58,8 +58,7 @@ by (induct t arbitrary: n rule: llam.induct) (simp_all add: perm_nat_def) -constdefs - freshen :: "llam \ name \ llam" +definition freshen :: "llam \ name \ llam" where "freshen t p \ vsub t 0 (lPar p)" lemma freshen_eqvt[eqvt]: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Nominal/Examples/SN.thy --- a/src/HOL/Nominal/Examples/SN.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Nominal/Examples/SN.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,5 @@ theory SN - imports Lam_Funs +imports Lam_Funs begin text {* Strong Normalisation proof from the Proofs and Types book *} @@ -158,8 +158,7 @@ subsection {* a fact about beta *} -constdefs - "NORMAL" :: "lam \ bool" +definition "NORMAL" :: "lam \ bool" where "NORMAL t \ \(\t'. t\\<^isub>\ t')" lemma NORMAL_Var: @@ -234,8 +233,7 @@ by (rule TrueI)+ text {* neutral terms *} -constdefs - NEUT :: "lam \ bool" +definition NEUT :: "lam \ bool" where "NEUT t \ (\a. t = Var a) \ (\t1 t2. t = App t1 t2)" (* a slight hack to get the first element of applications *) @@ -274,20 +272,19 @@ section {* Candidates *} -constdefs - "CR1" :: "ty \ bool" +definition "CR1" :: "ty \ bool" where "CR1 \ \ \t. (t\RED \ \ SN t)" - "CR2" :: "ty \ bool" +definition "CR2" :: "ty \ bool" where "CR2 \ \ \t t'. (t\RED \ \ t \\<^isub>\ t') \ t'\RED \" - "CR3_RED" :: "lam \ ty \ bool" +definition "CR3_RED" :: "lam \ ty \ bool" where "CR3_RED t \ \ \t'. t\\<^isub>\ t' \ t'\RED \" - "CR3" :: "ty \ bool" +definition "CR3" :: "ty \ bool" where "CR3 \ \ \t. (NEUT t \ CR3_RED t \) \ t\RED \" - "CR4" :: "ty \ bool" +definition "CR4" :: "ty \ bool" where "CR4 \ \ \t. (NEUT t \ NORMAL t) \t\RED \" lemma CR3_implies_CR4: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Nominal/Nominal.thy Mon Mar 01 13:42:31 2010 +0100 @@ -31,7 +31,7 @@ (* an auxiliary constant for the decision procedure involving *) (* permutations (to avoid loops when using perm-compositions) *) -constdefs +definition "perm_aux pi x \ pi\x" (* overloaded permutation operations *) @@ -187,20 +187,18 @@ section {* permutation equality *} (*==============================*) -constdefs - prm_eq :: "'x prm \ 'x prm \ bool" (" _ \ _ " [80,80] 80) +definition prm_eq :: "'x prm \ 'x prm \ bool" (" _ \ _ " [80,80] 80) where "pi1 \ pi2 \ \a::'x. pi1\a = pi2\a" section {* Support, Freshness and Supports*} (*========================================*) -constdefs - supp :: "'a \ ('x set)" +definition supp :: "'a \ ('x set)" where "supp x \ {a . (infinite {b . [(a,b)]\x \ x})}" - fresh :: "'x \ 'a \ bool" ("_ \ _" [80,80] 80) +definition fresh :: "'x \ 'a \ bool" ("_ \ _" [80,80] 80) where "a \ x \ a \ supp x" - supports :: "'x set \ 'a \ bool" (infixl "supports" 80) +definition supports :: "'x set \ 'a \ bool" (infixl "supports" 80) where "S supports x \ \a b. (a\S \ b\S \ [(a,b)]\x=x)" (* lemmas about supp *) @@ -400,14 +398,14 @@ (*=========================================================*) (* properties for being a permutation type *) -constdefs +definition "pt TYPE('a) TYPE('x) \ (\(x::'a). ([]::'x prm)\x = x) \ (\(pi1::'x prm) (pi2::'x prm) (x::'a). (pi1@pi2)\x = pi1\(pi2\x)) \ (\(pi1::'x prm) (pi2::'x prm) (x::'a). pi1 \ pi2 \ pi1\x = pi2\x)" (* properties for being an atom type *) -constdefs +definition "at TYPE('x) \ (\(x::'x). ([]::'x prm)\x = x) \ (\(a::'x) (b::'x) (pi::'x prm) (x::'x). ((a,b)#(pi::'x prm))\x = swap (a,b) (pi\x)) \ @@ -415,18 +413,18 @@ (infinite (UNIV::'x set))" (* property of two atom-types being disjoint *) -constdefs +definition "disjoint TYPE('x) TYPE('y) \ (\(pi::'x prm)(x::'y). pi\x = x) \ (\(pi::'y prm)(x::'x). pi\x = x)" (* composition property of two permutation on a type 'a *) -constdefs +definition "cp TYPE ('a) TYPE('x) TYPE('y) \ (\(pi2::'y prm) (pi1::'x prm) (x::'a) . pi1\(pi2\x) = (pi1\pi2)\(pi1\x))" (* property of having finite support *) -constdefs +definition "fs TYPE('a) TYPE('x) \ \(x::'a). finite ((supp x)::'x set)" section {* Lemmas about the atom-type properties*} @@ -2216,8 +2214,7 @@ section {* Facts about the support of finite sets of finitely supported things *} (*=============================================================================*) -constdefs - X_to_Un_supp :: "('a set) \ 'x set" +definition X_to_Un_supp :: "('a set) \ 'x set" where "X_to_Un_supp X \ \x\X. ((supp x)::'x set)" lemma UNION_f_eqvt: @@ -2838,8 +2835,7 @@ qed -- "packaging the freshness lemma into a function" -constdefs - fresh_fun :: "('x\'a)\'a" +definition fresh_fun :: "('x\'a)\'a" where "fresh_fun (h) \ THE fr. (\(a::'x). a\h \ (h a) = fr)" lemma fresh_fun_app: @@ -2970,8 +2966,7 @@ shows "pt TYPE('x\('a noption)) TYPE('x)" by (rule pt_fun_inst[OF at_pt_inst[OF at],OF pt_noption_inst[OF pt],OF at]) -constdefs - abs_fun :: "'x\'a\('x\('a noption))" ("[_]._" [100,100] 100) +definition abs_fun :: "'x\'a\('x\('a noption))" ("[_]._" [100,100] 100) where "[a].x \ (\b. (if b=a then nSome(x) else (if b\x then nSome([(a,b)]\x) else nNone)))" (* FIXME: should be called perm_if and placed close to the definition of permutations on bools *) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Number_Theory/MiscAlgebra.thy --- a/src/HOL/Number_Theory/MiscAlgebra.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Number_Theory/MiscAlgebra.thy Mon Mar 01 13:42:31 2010 +0100 @@ -31,8 +31,7 @@ *) -constdefs - units_of :: "('a, 'b) monoid_scheme => 'a monoid" +definition units_of :: "('a, 'b) monoid_scheme => 'a monoid" where "units_of G == (| carrier = Units G, Group.monoid.mult = Group.monoid.mult G, one = one G |)"; diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Number_Theory/Residues.thy --- a/src/HOL/Number_Theory/Residues.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Number_Theory/Residues.thy Mon Mar 01 13:42:31 2010 +0100 @@ -22,8 +22,7 @@ *) -constdefs - residue_ring :: "int => int ring" +definition residue_ring :: "int => int ring" where "residue_ring m == (| carrier = {0..m - 1}, mult = (%x y. (x * y) mod m), @@ -287,8 +286,7 @@ (* the definition of the phi function *) -constdefs - phi :: "int => nat" +definition phi :: "int => nat" where "phi m == card({ x. 0 < x & x < m & gcd x m = 1})" lemma phi_zero [simp]: "phi 0 = 0" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Number_Theory/UniqueFactorization.thy --- a/src/HOL/Number_Theory/UniqueFactorization.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Number_Theory/UniqueFactorization.thy Mon Mar 01 13:42:31 2010 +0100 @@ -67,8 +67,7 @@ "ALL i :# M. P i"? *) -constdefs - msetprod :: "('a => ('b::{power,comm_monoid_mult})) => 'a multiset => 'b" +definition msetprod :: "('a => ('b::{power,comm_monoid_mult})) => 'a multiset => 'b" where "msetprod f M == setprod (%x. (f x)^(count M x)) (set_of M)" syntax @@ -214,8 +213,7 @@ thus ?thesis by (simp add:multiset_eq_conv_count_eq) qed -constdefs - multiset_prime_factorization :: "nat => nat multiset" +definition multiset_prime_factorization :: "nat => nat multiset" where "multiset_prime_factorization n == if n > 0 then (THE M. ((ALL p : set_of M. prime p) & n = (PROD i :# M. i))) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Recdef.thy --- a/src/HOL/Recdef.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Recdef.thy Mon Mar 01 13:42:31 2010 +0100 @@ -27,15 +27,17 @@ wfrecI: "ALL z. (z, x) : R --> wfrec_rel R F z (g z) ==> wfrec_rel R F x (F g x)" -constdefs - cut :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b" +definition + cut :: "('a => 'b) => ('a * 'a)set => 'a => 'a => 'b" where "cut f r x == (%y. if (y,x):r then f y else undefined)" - adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool" +definition + adm_wf :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => bool" where "adm_wf R F == ALL f g x. (ALL z. (z, x) : R --> f z = g z) --> F f x = F g x" - wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b" +definition + wfrec :: "('a * 'a) set => (('a => 'b) => 'a => 'b) => 'a => 'b" where [code del]: "wfrec R F == %x. THE y. wfrec_rel R (%f x. F (cut f R x) x) x y" subsection{*Well-Founded Recursion*} diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/SET_Protocol/Message_SET.thy --- a/src/HOL/SET_Protocol/Message_SET.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/SET_Protocol/Message_SET.thy Mon Mar 01 13:42:31 2010 +0100 @@ -48,8 +48,7 @@ text{*The inverse of a symmetric key is itself; that of a public key is the private key and vice versa*} -constdefs - symKeys :: "key set" +definition symKeys :: "key set" where "symKeys == {K. invKey K = K}" text{*Agents. We allow any number of certification authorities, cardholders @@ -81,8 +80,7 @@ "{|x, y|}" == "CONST MPair x y" -constdefs - nat_of_agent :: "agent => nat" +definition nat_of_agent :: "agent => nat" where "nat_of_agent == agent_case (curry nat2_to_nat 0) (curry nat2_to_nat 1) (curry nat2_to_nat 2) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/SET_Protocol/Public_SET.thy --- a/src/HOL/SET_Protocol/Public_SET.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/SET_Protocol/Public_SET.thy Mon Mar 01 13:42:31 2010 +0100 @@ -161,21 +161,19 @@ subsection{*Encryption Primitives*} -constdefs - - EXcrypt :: "[key,key,msg,msg] => msg" +definition EXcrypt :: "[key,key,msg,msg] => msg" where --{*Extra Encryption*} (*K: the symmetric key EK: the public encryption key*) "EXcrypt K EK M m == {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m|}|}" - EXHcrypt :: "[key,key,msg,msg] => msg" +definition EXHcrypt :: "[key,key,msg,msg] => msg" where --{*Extra Encryption with Hashing*} (*K: the symmetric key EK: the public encryption key*) "EXHcrypt K EK M m == {|Crypt K {|M, Hash m|}, Crypt EK {|Key K, m, Hash M|}|}" - Enc :: "[key,key,key,msg] => msg" +definition Enc :: "[key,key,key,msg] => msg" where --{*Simple Encapsulation with SIGNATURE*} (*SK: the sender's signing key K: the symmetric key @@ -183,7 +181,7 @@ "Enc SK K EK M == {|Crypt K (sign SK M), Crypt EK (Key K)|}" - EncB :: "[key,key,key,msg,msg] => msg" +definition EncB :: "[key,key,key,msg,msg] => msg" where --{*Encapsulation with Baggage. Keys as above, and baggage b.*} "EncB SK K EK M b == {|Enc SK K EK {|M, Hash b|}, b|}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Set.thy --- a/src/HOL/Set.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Set.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1586,8 +1586,7 @@ subsubsection {* Inverse image of a function *} -constdefs - vimage :: "('a => 'b) => 'b set => 'a set" (infixr "-`" 90) +definition vimage :: "('a => 'b) => 'b set => 'a set" (infixr "-`" 90) where [code del]: "f -` B == {x. f x : B}" lemma vimage_eq [simp]: "(a : f -` B) = (f a : B)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Statespace/StateFun.thy --- a/src/HOL/Statespace/StateFun.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Statespace/StateFun.thy Mon Mar 01 13:42:31 2010 +0100 @@ -22,7 +22,7 @@ better compositionality, especially if you think of nested state spaces. *} -constdefs K_statefun:: "'a \ 'b \ 'a" "K_statefun c x \ c" +definition K_statefun :: "'a \ 'b \ 'a" where "K_statefun c x \ c" lemma K_statefun_apply [simp]: "K_statefun c x = c" by (simp add: K_statefun_def) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Comp.thy --- a/src/HOL/UNITY/Comp.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Comp.thy Mon Mar 01 13:42:31 2010 +0100 @@ -27,23 +27,20 @@ end -constdefs - component_of :: "'a program =>'a program=> bool" - (infixl "component'_of" 50) +definition component_of :: "'a program =>'a program=> bool" (infixl "component'_of" 50) where "F component_of H == \G. F ok G & F\G = H" - strict_component_of :: "'a program\'a program=> bool" - (infixl "strict'_component'_of" 50) +definition strict_component_of :: "'a program\'a program=> bool" (infixl "strict'_component'_of" 50) where "F strict_component_of H == F component_of H & F\H" - preserves :: "('a=>'b) => 'a program set" +definition preserves :: "('a=>'b) => 'a program set" where "preserves v == \z. stable {s. v s = z}" - localize :: "('a=>'b) => 'a program => 'a program" +definition localize :: "('a=>'b) => 'a program => 'a program" where "localize v F == mk_program(Init F, Acts F, AllowedActs F \ (\G \ preserves v. Acts G))" - funPair :: "['a => 'b, 'a => 'c, 'a] => 'b * 'c" +definition funPair :: "['a => 'b, 'a => 'c, 'a] => 'b * 'c" where "funPair f g == %x. (f x, g x)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Comp/AllocImpl.thy --- a/src/HOL/UNITY/Comp/AllocImpl.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Comp/AllocImpl.thy Mon Mar 01 13:42:31 2010 +0100 @@ -20,8 +20,7 @@ "'b merge" + dummy :: 'a (*dummy field for new variables*) -constdefs - non_dummy :: "('a,'b) merge_d => 'b merge" +definition non_dummy :: "('a,'b) merge_d => 'b merge" where "non_dummy s == (|In = In s, Out = Out s, iOut = iOut s|)" record 'b distr = diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Comp/Counter.thy --- a/src/HOL/UNITY/Comp/Counter.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Comp/Counter.thy Mon Mar 01 13:42:31 2010 +0100 @@ -16,25 +16,21 @@ datatype name = C | c nat types state = "name=>int" -consts - sum :: "[nat,state]=>int" - sumj :: "[nat, nat, state]=>int" - -primrec (* sum I s = sigma_{iint" where + (* sum I s = sigma_{iint" where "sumj 0 i s = 0" - "sumj (Suc n) i s = (if n=i then sum n s else s (c n) + sumj n i s)" +| "sumj (Suc n) i s = (if n=i then sum n s else s (c n) + sumj n i s)" types command = "(state*state)set" -constdefs - a :: "nat=>command" +definition a :: "nat=>command" where "a i == {(s, s'). s'=s(c i:= s (c i) + 1, C:= s C + 1)}" - Component :: "nat => state program" +definition Component :: "nat => state program" where "Component i == mk_total_program({s. s C = 0 & s (c i) = 0}, {a i}, \G \ preserves (%s. s (c i)). Acts G)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Comp/Counterc.thy --- a/src/HOL/UNITY/Comp/Counterc.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Comp/Counterc.thy Mon Mar 01 13:42:31 2010 +0100 @@ -21,25 +21,21 @@ C :: "state=>int" c :: "state=>nat=>int" -consts - sum :: "[nat,state]=>int" - sumj :: "[nat, nat, state]=>int" - -primrec (* sum I s = sigma_{iint" where + (* sum I s = sigma_{iint" where "sumj 0 i s = 0" - "sumj (Suc n) i s = (if n=i then sum n s else (c s) n + sumj n i s)" +| "sumj (Suc n) i s = (if n=i then sum n s else (c s) n + sumj n i s)" types command = "(state*state)set" -constdefs - a :: "nat=>command" +definition a :: "nat=>command" where "a i == {(s, s'). (c s') i = (c s) i + 1 & (C s') = (C s) + 1}" - Component :: "nat => state program" +definition Component :: "nat => state program" where "Component i == mk_total_program({s. C s = 0 & (c s) i = 0}, {a i}, \G \ preserves (%s. (c s) i). Acts G)" @@ -127,4 +123,4 @@ apply (auto intro!: sum0 p2_p3) done -end +end diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Comp/Priority.thy --- a/src/HOL/UNITY/Comp/Priority.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Comp/Priority.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/UNITY/Priority - ID: $Id$ Author: Sidi O Ehmety, Cambridge University Computer Laboratory Copyright 2001 University of Cambridge *) @@ -22,52 +21,50 @@ text{*Following the definitions given in section 4.4 *} -constdefs - highest :: "[vertex, (vertex*vertex)set]=>bool" +definition highest :: "[vertex, (vertex*vertex)set]=>bool" where "highest i r == A i r = {}" --{* i has highest priority in r *} - lowest :: "[vertex, (vertex*vertex)set]=>bool" +definition lowest :: "[vertex, (vertex*vertex)set]=>bool" where "lowest i r == R i r = {}" --{* i has lowest priority in r *} - act :: command +definition act :: command where "act i == {(s, s'). s'=reverse i s & highest i s}" - Component :: "vertex=>state program" +definition Component :: "vertex=>state program" where "Component i == mk_total_program({init}, {act i}, UNIV)" --{* All components start with the same initial state *} text{*Some Abbreviations *} -constdefs - Highest :: "vertex=>state set" +definition Highest :: "vertex=>state set" where "Highest i == {s. highest i s}" - Lowest :: "vertex=>state set" +definition Lowest :: "vertex=>state set" where "Lowest i == {s. lowest i s}" - Acyclic :: "state set" +definition Acyclic :: "state set" where "Acyclic == {s. acyclic s}" - Maximal :: "state set" +definition Maximal :: "state set" where --{* Every ``above'' set has a maximal vertex*} "Maximal == \i. {s. ~highest i s-->(\j \ above i s. highest j s)}" - Maximal' :: "state set" +definition Maximal' :: "state set" where --{* Maximal vertex: equivalent definition*} "Maximal' == \i. Highest i Un (\j. {s. j \ above i s} Int Highest j)" - Safety :: "state set" +definition Safety :: "state set" where "Safety == \i. {s. highest i s --> (\j \ neighbors i s. ~highest j s)}" (* Composition of a finite set of component; the vertex 'UNIV' is finite by assumption *) - system :: "state program" +definition system :: "state program" where "system == JN i. Component i" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Comp/PriorityAux.thy --- a/src/HOL/UNITY/Comp/PriorityAux.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Comp/PriorityAux.thy Mon Mar 01 13:42:31 2010 +0100 @@ -12,38 +12,37 @@ typedecl vertex -constdefs - symcl :: "(vertex*vertex)set=>(vertex*vertex)set" +definition symcl :: "(vertex*vertex)set=>(vertex*vertex)set" where "symcl r == r \ (r^-1)" --{* symmetric closure: removes the orientation of a relation*} - neighbors :: "[vertex, (vertex*vertex)set]=>vertex set" +definition neighbors :: "[vertex, (vertex*vertex)set]=>vertex set" where "neighbors i r == ((r \ r^-1)``{i}) - {i}" --{* Neighbors of a vertex i *} - R :: "[vertex, (vertex*vertex)set]=>vertex set" +definition R :: "[vertex, (vertex*vertex)set]=>vertex set" where "R i r == r``{i}" - A :: "[vertex, (vertex*vertex)set]=>vertex set" +definition A :: "[vertex, (vertex*vertex)set]=>vertex set" where "A i r == (r^-1)``{i}" - reach :: "[vertex, (vertex*vertex)set]=> vertex set" +definition reach :: "[vertex, (vertex*vertex)set]=> vertex set" where "reach i r == (r^+)``{i}" --{* reachable and above vertices: the original notation was R* and A* *} - above :: "[vertex, (vertex*vertex)set]=> vertex set" +definition above :: "[vertex, (vertex*vertex)set]=> vertex set" where "above i r == ((r^-1)^+)``{i}" - reverse :: "[vertex, (vertex*vertex) set]=>(vertex*vertex)set" +definition reverse :: "[vertex, (vertex*vertex) set]=>(vertex*vertex)set" where "reverse i r == (r - {(x,y). x=i | y=i} \ r) \ ({(x,y). x=i|y=i} \ r)^-1" - derive1 :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" +definition derive1 :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" where --{* The original definition *} "derive1 i r q == symcl r = symcl q & (\k k'. k\i & k'\i -->((k,k'):r) = ((k,k'):q)) & A i r = {} & R i q = {}" - derive :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" +definition derive :: "[vertex, (vertex*vertex)set, (vertex*vertex)set]=>bool" where --{* Our alternative definition *} "derive i r q == A i r = {} & (q = reverse i r)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Comp/Progress.thy --- a/src/HOL/UNITY/Comp/Progress.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Comp/Progress.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/UNITY/Progress - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 2003 University of Cambridge @@ -13,11 +12,10 @@ subsection {*The Composition of Two Single-Assignment Programs*} text{*Thesis Section 4.4.2*} -constdefs - FF :: "int program" +definition FF :: "int program" where "FF == mk_total_program (UNIV, {range (\x. (x, x+1))}, UNIV)" - GG :: "int program" +definition GG :: "int program" where "GG == mk_total_program (UNIV, {range (\x. (x, 2*x))}, UNIV)" subsubsection {*Calculating @{term "wens_set FF (atLeast k)"}*} diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Comp/TimerArray.thy --- a/src/HOL/UNITY/Comp/TimerArray.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Comp/TimerArray.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/UNITY/TimerArray.thy - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge @@ -10,14 +9,13 @@ types 'a state = "nat * 'a" (*second component allows new variables*) -constdefs - count :: "'a state => nat" +definition count :: "'a state => nat" where "count s == fst s" - decr :: "('a state * 'a state) set" +definition decr :: "('a state * 'a state) set" where "decr == UN n uu. {((Suc n, uu), (n,uu))}" - Timer :: "'a state program" +definition Timer :: "'a state program" where "Timer == mk_total_program (UNIV, {decr}, UNIV)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Constrains.thy --- a/src/HOL/UNITY/Constrains.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Constrains.thy Mon Mar 01 13:42:31 2010 +0100 @@ -32,22 +32,21 @@ | Acts: "[| act: Acts F; s \ reachable F; (s,s'): act |] ==> s' \ reachable F" -constdefs - Constrains :: "['a set, 'a set] => 'a program set" (infixl "Co" 60) +definition Constrains :: "['a set, 'a set] => 'a program set" (infixl "Co" 60) where "A Co B == {F. F \ (reachable F \ A) co B}" - Unless :: "['a set, 'a set] => 'a program set" (infixl "Unless" 60) +definition Unless :: "['a set, 'a set] => 'a program set" (infixl "Unless" 60) where "A Unless B == (A-B) Co (A \ B)" - Stable :: "'a set => 'a program set" +definition Stable :: "'a set => 'a program set" where "Stable A == A Co A" (*Always is the weak form of "invariant"*) - Always :: "'a set => 'a program set" +definition Always :: "'a set => 'a program set" where "Always A == {F. Init F \ A} \ Stable A" (*Polymorphic in both states and the meaning of \ *) - Increasing :: "['a => 'b::{order}] => 'a program set" +definition Increasing :: "['a => 'b::{order}] => 'a program set" where "Increasing f == \z. Stable {s. z \ f s}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/FP.thy --- a/src/HOL/UNITY/FP.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/FP.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/UNITY/FP - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge @@ -10,12 +9,10 @@ theory FP imports UNITY begin -constdefs - - FP_Orig :: "'a program => 'a set" +definition FP_Orig :: "'a program => 'a set" where "FP_Orig F == Union{A. ALL B. F : stable (A Int B)}" - FP :: "'a program => 'a set" +definition FP :: "'a program => 'a set" where "FP F == {s. F : stable {s}}" lemma stable_FP_Orig_Int: "F : stable (FP_Orig F Int B)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Follows.thy --- a/src/HOL/UNITY/Follows.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Follows.thy Mon Mar 01 13:42:31 2010 +0100 @@ -7,10 +7,7 @@ theory Follows imports SubstAx ListOrder Multiset begin -constdefs - - Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set" - (infixl "Fols" 65) +definition Follows :: "['a => 'b::{order}, 'a => 'b::{order}] => 'a program set" (infixl "Fols" 65) where "f Fols g == Increasing g \ Increasing f Int Always {s. f s \ g s} Int (\k. {s. k \ g s} LeadsTo {s. k \ f s})" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Guar.thy --- a/src/HOL/UNITY/Guar.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Guar.thy Mon Mar 01 13:42:31 2010 +0100 @@ -22,51 +22,47 @@ text{*Existential and Universal properties. I formalize the two-program case, proving equivalence with Chandy and Sanders's n-ary definitions*} -constdefs - - ex_prop :: "'a program set => bool" +definition ex_prop :: "'a program set => bool" where "ex_prop X == \F G. F ok G -->F \ X | G \ X --> (F\G) \ X" - strict_ex_prop :: "'a program set => bool" +definition strict_ex_prop :: "'a program set => bool" where "strict_ex_prop X == \F G. F ok G --> (F \ X | G \ X) = (F\G \ X)" - uv_prop :: "'a program set => bool" +definition uv_prop :: "'a program set => bool" where "uv_prop X == SKIP \ X & (\F G. F ok G --> F \ X & G \ X --> (F\G) \ X)" - strict_uv_prop :: "'a program set => bool" +definition strict_uv_prop :: "'a program set => bool" where "strict_uv_prop X == SKIP \ X & (\F G. F ok G --> (F \ X & G \ X) = (F\G \ X))" text{*Guarantees properties*} -constdefs - - guar :: "['a program set, 'a program set] => 'a program set" - (infixl "guarantees" 55) (*higher than membership, lower than Co*) +definition guar :: "['a program set, 'a program set] => 'a program set" (infixl "guarantees" 55) where + (*higher than membership, lower than Co*) "X guarantees Y == {F. \G. F ok G --> F\G \ X --> F\G \ Y}" (* Weakest guarantees *) - wg :: "['a program, 'a program set] => 'a program set" +definition wg :: "['a program, 'a program set] => 'a program set" where "wg F Y == Union({X. F \ X guarantees Y})" (* Weakest existential property stronger than X *) - wx :: "('a program) set => ('a program)set" +definition wx :: "('a program) set => ('a program)set" where "wx X == Union({Y. Y \ X & ex_prop Y})" (*Ill-defined programs can arise through "Join"*) - welldef :: "'a program set" +definition welldef :: "'a program set" where "welldef == {F. Init F \ {}}" - refines :: "['a program, 'a program, 'a program set] => bool" - ("(3_ refines _ wrt _)" [10,10,10] 10) +definition refines :: "['a program, 'a program, 'a program set] => bool" + ("(3_ refines _ wrt _)" [10,10,10] 10) where "G refines F wrt X == \H. (F ok H & G ok H & F\H \ welldef \ X) --> (G\H \ welldef \ X)" - iso_refines :: "['a program, 'a program, 'a program set] => bool" - ("(3_ iso'_refines _ wrt _)" [10,10,10] 10) +definition iso_refines :: "['a program, 'a program, 'a program set] => bool" + ("(3_ iso'_refines _ wrt _)" [10,10,10] 10) where "G iso_refines F wrt X == F \ welldef \ X --> G \ welldef \ X" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Lift_prog.thy --- a/src/HOL/UNITY/Lift_prog.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Lift_prog.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/UNITY/Lift_prog.thy - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1999 University of Cambridge @@ -10,30 +9,28 @@ theory Lift_prog imports Rename begin -constdefs - - insert_map :: "[nat, 'b, nat=>'b] => (nat=>'b)" +definition insert_map :: "[nat, 'b, nat=>'b] => (nat=>'b)" where "insert_map i z f k == if k'b] => (nat=>'b)" +definition delete_map :: "[nat, nat=>'b] => (nat=>'b)" where "delete_map i g k == if k'b) * 'c)] => (nat=>'b) * 'c" +definition lift_map :: "[nat, 'b * ((nat=>'b) * 'c)] => (nat=>'b) * 'c" where "lift_map i == %(s,(f,uu)). (insert_map i s f, uu)" - drop_map :: "[nat, (nat=>'b) * 'c] => 'b * ((nat=>'b) * 'c)" +definition drop_map :: "[nat, (nat=>'b) * 'c] => 'b * ((nat=>'b) * 'c)" where "drop_map i == %(g, uu). (g i, (delete_map i g, uu))" - lift_set :: "[nat, ('b * ((nat=>'b) * 'c)) set] => ((nat=>'b) * 'c) set" +definition lift_set :: "[nat, ('b * ((nat=>'b) * 'c)) set] => ((nat=>'b) * 'c) set" where "lift_set i A == lift_map i ` A" - lift :: "[nat, ('b * ((nat=>'b) * 'c)) program] => ((nat=>'b) * 'c) program" +definition lift :: "[nat, ('b * ((nat=>'b) * 'c)) program] => ((nat=>'b) * 'c) program" where "lift i == rename (lift_map i)" (*simplifies the expression of specifications*) - sub :: "['a, 'a=>'b] => 'b" +definition sub :: "['a, 'a=>'b] => 'b" where "sub == %i f. f i" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/ListOrder.thy --- a/src/HOL/UNITY/ListOrder.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/ListOrder.thy Mon Mar 01 13:42:31 2010 +0100 @@ -42,11 +42,10 @@ end -constdefs - Le :: "(nat*nat) set" +definition Le :: "(nat*nat) set" where "Le == {(x,y). x <= y}" - Ge :: "(nat*nat) set" +definition Ge :: "(nat*nat) set" where "Ge == {(x,y). y <= x}" abbreviation diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/PPROD.thy --- a/src/HOL/UNITY/PPROD.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/PPROD.thy Mon Mar 01 13:42:31 2010 +0100 @@ -10,9 +10,8 @@ theory PPROD imports Lift_prog begin -constdefs - PLam :: "[nat set, nat => ('b * ((nat=>'b) * 'c)) program] - => ((nat=>'b) * 'c) program" +definition PLam :: "[nat set, nat => ('b * ((nat=>'b) * 'c)) program] + => ((nat=>'b) * 'c) program" where "PLam I F == \i \ I. lift i (F i)" syntax diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/ProgressSets.thy --- a/src/HOL/UNITY/ProgressSets.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/ProgressSets.thy Mon Mar 01 13:42:31 2010 +0100 @@ -19,13 +19,12 @@ subsection {*Complete Lattices and the Operator @{term cl}*} -constdefs - lattice :: "'a set set => bool" +definition lattice :: "'a set set => bool" where --{*Meier calls them closure sets, but they are just complete lattices*} "lattice L == (\M. M \ L --> \M \ L) & (\M. M \ L --> \M \ L)" - cl :: "['a set set, 'a set] => 'a set" +definition cl :: "['a set set, 'a set] => 'a set" where --{*short for ``closure''*} "cl L r == \{x. x\L & r \ x}" @@ -143,12 +142,11 @@ text{*A progress set satisfies certain closure conditions and is a simple way of including the set @{term "wens_set F B"}.*} -constdefs - closed :: "['a program, 'a set, 'a set, 'a set set] => bool" +definition closed :: "['a program, 'a set, 'a set, 'a set set] => bool" where "closed F T B L == \M. \act \ Acts F. B\M & T\M \ L --> T \ (B \ wp act M) \ L" - progress_set :: "['a program, 'a set, 'a set] => 'a set set set" +definition progress_set :: "['a program, 'a set, 'a set] => 'a set set set" where "progress_set F T B == {L. lattice L & B \ L & T \ L & closed F T B L}" @@ -324,12 +322,11 @@ subsubsection {*Lattices and Relations*} text{*From Meier's thesis, section 4.5.3*} -constdefs - relcl :: "'a set set => ('a * 'a) set" +definition relcl :: "'a set set => ('a * 'a) set" where -- {*Derived relation from a lattice*} "relcl L == {(x,y). y \ cl L {x}}" - latticeof :: "('a * 'a) set => 'a set set" +definition latticeof :: "('a * 'a) set => 'a set set" where -- {*Derived lattice from a relation: the set of upwards-closed sets*} "latticeof r == {X. \s t. s \ X & (s,t) \ r --> t \ X}" @@ -405,8 +402,7 @@ subsubsection {*Decoupling Theorems*} -constdefs - decoupled :: "['a program, 'a program] => bool" +definition decoupled :: "['a program, 'a program] => bool" where "decoupled F G == \act \ Acts F. \B. G \ stable B --> G \ stable (wp act B)" @@ -469,8 +465,7 @@ subsection{*Composition Theorems Based on Monotonicity and Commutativity*} subsubsection{*Commutativity of @{term "cl L"} and assignment.*} -constdefs - commutes :: "['a program, 'a set, 'a set, 'a set set] => bool" +definition commutes :: "['a program, 'a set, 'a set, 'a set set] => bool" where "commutes F T B L == \M. \act \ Acts F. B \ M --> cl L (T \ wp act M) \ T \ (B \ wp act (cl L (T\M)))" @@ -511,8 +506,7 @@ text{*Possibly move to Relation.thy, after @{term single_valued}*} -constdefs - funof :: "[('a*'b)set, 'a] => 'b" +definition funof :: "[('a*'b)set, 'a] => 'b" where "funof r == (\x. THE y. (x,y) \ r)" lemma funof_eq: "[|single_valued r; (x,y) \ r|] ==> funof r x = y" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Project.thy --- a/src/HOL/UNITY/Project.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Project.thy Mon Mar 01 13:42:31 2010 +0100 @@ -11,19 +11,18 @@ theory Project imports Extend begin -constdefs - projecting :: "['c program => 'c set, 'a*'b => 'c, - 'a program, 'c program set, 'a program set] => bool" +definition projecting :: "['c program => 'c set, 'a*'b => 'c, + 'a program, 'c program set, 'a program set] => bool" where "projecting C h F X' X == \G. extend h F\G \ X' --> F\project h (C G) G \ X" - extending :: "['c program => 'c set, 'a*'b => 'c, 'a program, - 'c program set, 'a program set] => bool" +definition extending :: "['c program => 'c set, 'a*'b => 'c, 'a program, + 'c program set, 'a program set] => bool" where "extending C h F Y' Y == \G. extend h F ok G --> F\project h (C G) G \ Y --> extend h F\G \ Y'" - subset_closed :: "'a set set => bool" +definition subset_closed :: "'a set set => bool" where "subset_closed U == \A \ U. Pow A \ U" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Rename.thy --- a/src/HOL/UNITY/Rename.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Rename.thy Mon Mar 01 13:42:31 2010 +0100 @@ -8,9 +8,7 @@ theory Rename imports Extend begin -constdefs - - rename :: "['a => 'b, 'a program] => 'b program" +definition rename :: "['a => 'b, 'a program] => 'b program" where "rename h == extend (%(x,u::unit). h x)" declare image_inv_f_f [simp] image_surj_f_inv_f [simp] diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Simple/Channel.thy --- a/src/HOL/UNITY/Simple/Channel.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Simple/Channel.thy Mon Mar 01 13:42:31 2010 +0100 @@ -15,8 +15,7 @@ consts F :: "state program" -constdefs - minSet :: "nat set => nat option" +definition minSet :: "nat set => nat option" where "minSet A == if A={} then None else Some (LEAST x. x \ A)" axioms diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Simple/Common.thy --- a/src/HOL/UNITY/Simple/Common.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Simple/Common.thy Mon Mar 01 13:42:31 2010 +0100 @@ -24,11 +24,10 @@ fasc: "m \ ftime n" gasc: "m \ gtime n" -constdefs - common :: "nat set" +definition common :: "nat set" where "common == {n. ftime n = n & gtime n = n}" - maxfg :: "nat => nat set" +definition maxfg :: "nat => nat set" where "maxfg m == {t. t \ max (ftime m) (gtime m)}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Simple/NSP_Bad.thy --- a/src/HOL/UNITY/Simple/NSP_Bad.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Simple/NSP_Bad.thy Mon Mar 01 13:42:31 2010 +0100 @@ -53,8 +53,7 @@ & Says B' A3 (Crypt (pubK A3) {|Nonce NA, Nonce NB|}) \ set s3}" -constdefs - Nprg :: "state program" +definition Nprg :: "state program" where (*Initial trace is empty*) "Nprg == mk_total_program({[]}, {Fake, NS1, NS2, NS3}, UNIV)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Simple/Reach.thy --- a/src/HOL/UNITY/Simple/Reach.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Simple/Reach.thy Mon Mar 01 13:42:31 2010 +0100 @@ -17,21 +17,19 @@ edges :: "(vertex*vertex) set" -constdefs - - asgt :: "[vertex,vertex] => (state*state) set" +definition asgt :: "[vertex,vertex] => (state*state) set" where "asgt u v == {(s,s'). s' = s(v:= s u | s v)}" - Rprg :: "state program" +definition Rprg :: "state program" where "Rprg == mk_total_program ({%v. v=init}, \(u,v)\edges. {asgt u v}, UNIV)" - reach_invariant :: "state set" +definition reach_invariant :: "state set" where "reach_invariant == {s. (\v. s v --> (init, v) \ edges^*) & s init}" - fixedpoint :: "state set" +definition fixedpoint :: "state set" where "fixedpoint == {s. \(u,v)\edges. s u --> s v}" - metric :: "state => nat" +definition metric :: "state => nat" where "metric s == card {v. ~ s v}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Simple/Reachability.thy --- a/src/HOL/UNITY/Simple/Reachability.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Simple/Reachability.thy Mon Mar 01 13:42:31 2010 +0100 @@ -25,23 +25,22 @@ base: "v \ V ==> ((v,v) \ REACHABLE)" | step: "((u,v) \ REACHABLE) & (v,w) \ E ==> ((u,w) \ REACHABLE)" -constdefs - reachable :: "vertex => state set" +definition reachable :: "vertex => state set" where "reachable p == {s. reach s p}" - nmsg_eq :: "nat => edge => state set" +definition nmsg_eq :: "nat => edge => state set" where "nmsg_eq k == %e. {s. nmsg s e = k}" - nmsg_gt :: "nat => edge => state set" +definition nmsg_gt :: "nat => edge => state set" where "nmsg_gt k == %e. {s. k < nmsg s e}" - nmsg_gte :: "nat => edge => state set" +definition nmsg_gte :: "nat => edge => state set" where "nmsg_gte k == %e. {s. k \ nmsg s e}" - nmsg_lte :: "nat => edge => state set" +definition nmsg_lte :: "nat => edge => state set" where "nmsg_lte k == %e. {s. nmsg s e \ k}" - final :: "state set" +definition final :: "state set" where "final == (\v\V. reachable v <==> {s. (root, v) \ REACHABLE}) \ (INTER E (nmsg_eq 0))" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Simple/Token.thy --- a/src/HOL/UNITY/Simple/Token.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Simple/Token.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/UNITY/Token - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge *) @@ -24,17 +23,16 @@ proc :: "nat => pstate" -constdefs - HasTok :: "nat => state set" +definition HasTok :: "nat => state set" where "HasTok i == {s. token s = i}" - H :: "nat => state set" +definition H :: "nat => state set" where "H i == {s. proc s i = Hungry}" - E :: "nat => state set" +definition E :: "nat => state set" where "E i == {s. proc s i = Eating}" - T :: "nat => state set" +definition T :: "nat => state set" where "T i == {s. proc s i = Thinking}" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/SubstAx.thy --- a/src/HOL/UNITY/SubstAx.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/SubstAx.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/UNITY/SubstAx - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge @@ -10,11 +9,10 @@ theory SubstAx imports WFair Constrains begin -constdefs - Ensures :: "['a set, 'a set] => 'a program set" (infixl "Ensures" 60) +definition Ensures :: "['a set, 'a set] => 'a program set" (infixl "Ensures" 60) where "A Ensures B == {F. F \ (reachable F \ A) ensures B}" - LeadsTo :: "['a set, 'a set] => 'a program set" (infixl "LeadsTo" 60) +definition LeadsTo :: "['a set, 'a set] => 'a program set" (infixl "LeadsTo" 60) where "A LeadsTo B == {F. F \ (reachable F \ A) leadsTo B}" notation (xsymbols) diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/Transformers.thy --- a/src/HOL/UNITY/Transformers.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/Transformers.thy Mon Mar 01 13:42:31 2010 +0100 @@ -20,16 +20,15 @@ subsection{*Defining the Predicate Transformers @{term wp}, @{term awp} and @{term wens}*} -constdefs - wp :: "[('a*'a) set, 'a set] => 'a set" +definition wp :: "[('a*'a) set, 'a set] => 'a set" where --{*Dijkstra's weakest-precondition operator (for an individual command)*} "wp act B == - (act^-1 `` (-B))" - awp :: "['a program, 'a set] => 'a set" +definition awp :: "['a program, 'a set] => 'a set" where --{*Dijkstra's weakest-precondition operator (for a program)*} "awp F B == (\act \ Acts F. wp act B)" - wens :: "['a program, ('a*'a) set, 'a set] => 'a set" +definition wens :: "['a program, ('a*'a) set, 'a set] => 'a set" where --{*The weakest-ensures transformer*} "wens F act B == gfp(\X. (wp act B \ awp F (B \ X)) \ B)" @@ -335,11 +334,10 @@ text{*Next, we express the @{term "wens_set"} for single-assignment programs*} -constdefs - wens_single_finite :: "[('a*'a) set, 'a set, nat] => 'a set" +definition wens_single_finite :: "[('a*'a) set, 'a set, nat] => 'a set" where "wens_single_finite act B k == \i \ atMost k. (wp act ^^ i) B" - wens_single :: "[('a*'a) set, 'a set] => 'a set" +definition wens_single :: "[('a*'a) set, 'a set] => 'a set" where "wens_single act B == \i. (wp act ^^ i) B" lemma wens_single_Un_eq: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/UNITY.thy --- a/src/HOL/UNITY/UNITY.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/UNITY.thy Mon Mar 01 13:42:31 2010 +0100 @@ -17,40 +17,39 @@ allowed :: ('a * 'a)set set). Id \ acts & Id: allowed}" by blast -constdefs - Acts :: "'a program => ('a * 'a)set set" +definition Acts :: "'a program => ('a * 'a)set set" where "Acts F == (%(init, acts, allowed). acts) (Rep_Program F)" - "constrains" :: "['a set, 'a set] => 'a program set" (infixl "co" 60) +definition "constrains" :: "['a set, 'a set] => 'a program set" (infixl "co" 60) where "A co B == {F. \act \ Acts F. act``A \ B}" - unless :: "['a set, 'a set] => 'a program set" (infixl "unless" 60) +definition unless :: "['a set, 'a set] => 'a program set" (infixl "unless" 60) where "A unless B == (A-B) co (A \ B)" - mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set) - => 'a program" +definition mk_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set) + => 'a program" where "mk_program == %(init, acts, allowed). Abs_Program (init, insert Id acts, insert Id allowed)" - Init :: "'a program => 'a set" +definition Init :: "'a program => 'a set" where "Init F == (%(init, acts, allowed). init) (Rep_Program F)" - AllowedActs :: "'a program => ('a * 'a)set set" +definition AllowedActs :: "'a program => ('a * 'a)set set" where "AllowedActs F == (%(init, acts, allowed). allowed) (Rep_Program F)" - Allowed :: "'a program => 'a program set" +definition Allowed :: "'a program => 'a program set" where "Allowed F == {G. Acts G \ AllowedActs F}" - stable :: "'a set => 'a program set" +definition stable :: "'a set => 'a program set" where "stable A == A co A" - strongest_rhs :: "['a program, 'a set] => 'a set" +definition strongest_rhs :: "['a program, 'a set] => 'a set" where "strongest_rhs F A == Inter {B. F \ A co B}" - invariant :: "'a set => 'a program set" +definition invariant :: "'a set => 'a program set" where "invariant A == {F. Init F \ A} \ stable A" - increasing :: "['a => 'b::{order}] => 'a program set" +definition increasing :: "['a => 'b::{order}] => 'a program set" where --{*Polymorphic in both states and the meaning of @{text "\"}*} "increasing f == \z. stable {s. z \ f s}" @@ -357,20 +356,19 @@ subsection{*Partial versus Total Transitions*} -constdefs - totalize_act :: "('a * 'a)set => ('a * 'a)set" +definition totalize_act :: "('a * 'a)set => ('a * 'a)set" where "totalize_act act == act \ Id_on (-(Domain act))" - totalize :: "'a program => 'a program" +definition totalize :: "'a program => 'a program" where "totalize F == mk_program (Init F, totalize_act ` Acts F, AllowedActs F)" - mk_total_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set) - => 'a program" +definition mk_total_program :: "('a set * ('a * 'a)set set * ('a * 'a)set set) + => 'a program" where "mk_total_program args == totalize (mk_program args)" - all_total :: "'a program => bool" +definition all_total :: "'a program => bool" where "all_total F == \act \ Acts F. Domain act = UNIV" lemma insert_Id_image_Acts: "f Id = Id ==> insert Id (f`Acts F) = f ` Acts F" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/UNITY/WFair.thy --- a/src/HOL/UNITY/WFair.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/UNITY/WFair.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/UNITY/WFair - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge @@ -33,16 +32,17 @@ is the impossibility law for leads-to. *} -constdefs +definition --{*This definition specifies conditional fairness. The rest of the theory is generic to all forms of fairness. To get weak fairness, conjoin the inclusion below with @{term "A \ Domain act"}, which specifies that the action is enabled over all of @{term A}.*} - transient :: "'a set => 'a program set" + transient :: "'a set => 'a program set" where "transient A == {F. \act\Acts F. act``A \ -A}" - ensures :: "['a set, 'a set] => 'a program set" (infixl "ensures" 60) +definition + ensures :: "['a set, 'a set] => 'a program set" (infixl "ensures" 60) where "A ensures B == (A-B co A \ B) \ transient (A-B)" @@ -59,13 +59,11 @@ | Union: "\A \ S. (A,B) \ leads F ==> (Union S, B) \ leads F" -constdefs - - leadsTo :: "['a set, 'a set] => 'a program set" (infixl "leadsTo" 60) +definition leadsTo :: "['a set, 'a set] => 'a program set" (infixl "leadsTo" 60) where --{*visible version of the LEADS-TO relation*} "A leadsTo B == {F. (A,B) \ leads F}" - wlt :: "['a program, 'a set] => 'a set" +definition wlt :: "['a program, 'a set] => 'a set" where --{*predicate transformer: the largest set that leads to @{term B}*} "wlt F B == Union {A. F \ A leadsTo B}" @@ -641,4 +639,4 @@ apply blast+ done -end +end \ No newline at end of file diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Word/WordDefinition.thy --- a/src/HOL/Word/WordDefinition.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Word/WordDefinition.thy Mon Mar 01 13:42:31 2010 +0100 @@ -165,14 +165,13 @@ where "word_pred a = word_of_int (Int.pred (uint a))" -constdefs - udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50) +definition udvd :: "'a::len word => 'a::len word => bool" (infixl "udvd" 50) where "a udvd b == EX n>=0. uint b = n * uint a" - word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50) +definition word_sle :: "'a :: len word => 'a word => bool" ("(_/ <=s _)" [50, 51] 50) where "a <=s b == sint a <= sint b" - word_sless :: "'a :: len word => 'a word => bool" ("(_/ 'a word => bool" ("(_/ nat => 'a word" +definition setBit :: "'a :: len0 word => nat => 'a word" where "setBit w n == set_bit w n True" - clearBit :: "'a :: len0 word => nat => 'a word" +definition clearBit :: "'a :: len0 word => nat => 'a word" where "clearBit w n == set_bit w n False" subsection "Shift operations" -constdefs - sshiftr1 :: "'a :: len word => 'a word" +definition sshiftr1 :: "'a :: len word => 'a word" where "sshiftr1 w == word_of_int (bin_rest (sint w))" - bshiftr1 :: "bool => 'a :: len word => 'a word" +definition bshiftr1 :: "bool => 'a :: len word => 'a word" where "bshiftr1 b w == of_bl (b # butlast (to_bl w))" - sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55) +definition sshiftr :: "'a :: len word => nat => 'a word" (infixl ">>>" 55) where "w >>> n == (sshiftr1 ^^ n) w" - mask :: "nat => 'a::len word" +definition mask :: "nat => 'a::len word" where "mask n == (1 << n) - 1" - revcast :: "'a :: len0 word => 'b :: len0 word" +definition revcast :: "'a :: len0 word => 'b :: len0 word" where "revcast w == of_bl (takefill False (len_of TYPE('b)) (to_bl w))" - slice1 :: "nat => 'a :: len0 word => 'b :: len0 word" +definition slice1 :: "nat => 'a :: len0 word => 'b :: len0 word" where "slice1 n w == of_bl (takefill False n (to_bl w))" - slice :: "nat => 'a :: len0 word => 'b :: len0 word" +definition slice :: "nat => 'a :: len0 word => 'b :: len0 word" where "slice n w == slice1 (size w - n) w" subsection "Rotation" -constdefs - rotater1 :: "'a list => 'a list" +definition rotater1 :: "'a list => 'a list" where "rotater1 ys == case ys of [] => [] | x # xs => last ys # butlast ys" - rotater :: "nat => 'a list => 'a list" +definition rotater :: "nat => 'a list => 'a list" where "rotater n == rotater1 ^^ n" - word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word" +definition word_rotr :: "nat => 'a :: len0 word => 'a :: len0 word" where "word_rotr n w == of_bl (rotater n (to_bl w))" - word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word" +definition word_rotl :: "nat => 'a :: len0 word => 'a :: len0 word" where "word_rotl n w == of_bl (rotate n (to_bl w))" - word_roti :: "int => 'a :: len0 word => 'a :: len0 word" +definition word_roti :: "int => 'a :: len0 word => 'a :: len0 word" where "word_roti i w == if i >= 0 then word_rotr (nat i) w else word_rotl (nat (- i)) w" subsection "Split and cat operations" -constdefs - word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word" +definition word_cat :: "'a :: len0 word => 'b :: len0 word => 'c :: len0 word" where "word_cat a b == word_of_int (bin_cat (uint a) (len_of TYPE ('b)) (uint b))" - word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)" +definition word_split :: "'a :: len0 word => ('b :: len0 word) * ('c :: len0 word)" where "word_split a == case bin_split (len_of TYPE ('c)) (uint a) of (u, v) => (word_of_int u, word_of_int v)" - word_rcat :: "'a :: len0 word list => 'b :: len0 word" +definition word_rcat :: "'a :: len0 word list => 'b :: len0 word" where "word_rcat ws == word_of_int (bin_rcat (len_of TYPE ('a)) (map uint ws))" - word_rsplit :: "'a :: len0 word => 'b :: len word list" +definition word_rsplit :: "'a :: len0 word => 'b :: len word list" where "word_rsplit w == map word_of_int (bin_rsplit (len_of TYPE ('b)) (len_of TYPE ('a), uint w))" -constdefs - -- "Largest representable machine integer." - max_word :: "'a::len word" +definition max_word :: "'a::len word" -- "Largest representable machine integer." where "max_word \ word_of_int (2 ^ len_of TYPE('a) - 1)" -consts - of_bool :: "bool \ 'a::len word" -primrec +primrec of_bool :: "bool \ 'a::len word" where "of_bool False = 0" - "of_bool True = 1" +| "of_bool True = 1" lemmas of_nth_def = word_set_bits_def diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/Word/WordGenLib.thy --- a/src/HOL/Word/WordGenLib.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/Word/WordGenLib.thy Mon Mar 01 13:42:31 2010 +0100 @@ -344,8 +344,7 @@ apply (case_tac "1+n = 0", auto) done -constdefs - word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" +definition word_rec :: "'a \ ('b::len word \ 'a \ 'a) \ 'b word \ 'a" where "word_rec forZero forSuc n \ nat_rec forZero (forSuc \ of_nat) (unat n)" lemma word_rec_0: "word_rec z s 0 = z" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/ZF/Games.thy --- a/src/HOL/ZF/Games.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/ZF/Games.thy Mon Mar 01 13:42:31 2010 +0100 @@ -9,12 +9,13 @@ imports MainZF begin -constdefs - fixgames :: "ZF set \ ZF set" +definition fixgames :: "ZF set \ ZF set" where "fixgames A \ { Opair l r | l r. explode l \ A & explode r \ A}" - games_lfp :: "ZF set" + +definition games_lfp :: "ZF set" where "games_lfp \ lfp fixgames" - games_gfp :: "ZF set" + +definition games_gfp :: "ZF set" where "games_gfp \ gfp fixgames" lemma mono_fixgames: "mono (fixgames)" @@ -42,12 +43,13 @@ by auto qed -constdefs - left_option :: "ZF \ ZF \ bool" +definition left_option :: "ZF \ ZF \ bool" where "left_option g opt \ (Elem opt (Fst g))" - right_option :: "ZF \ ZF \ bool" + +definition right_option :: "ZF \ ZF \ bool" where "right_option g opt \ (Elem opt (Snd g))" - is_option_of :: "(ZF * ZF) set" + +definition is_option_of :: "(ZF * ZF) set" where "is_option_of \ { (opt, g) | opt g. g \ games_gfp \ (left_option g opt \ right_option g opt) }" lemma games_lfp_subset_gfp: "games_lfp \ games_gfp" @@ -190,14 +192,16 @@ typedef game = games_lfp by (blast intro: games_lfp_nonempty) -constdefs - left_options :: "game \ game zet" +definition left_options :: "game \ game zet" where "left_options g \ zimage Abs_game (zexplode (Fst (Rep_game g)))" - right_options :: "game \ game zet" + +definition right_options :: "game \ game zet" where "right_options g \ zimage Abs_game (zexplode (Snd (Rep_game g)))" - options :: "game \ game zet" + +definition options :: "game \ game zet" where "options g \ zunion (left_options g) (right_options g)" - Game :: "game zet \ game zet \ game" + +definition Game :: "game zet \ game zet \ game" where "Game L R \ Abs_game (Opair (zimplode (zimage Rep_game L)) (zimplode (zimage Rep_game R)))" lemma Repl_Rep_game_Abs_game: "\ e. Elem e z \ e \ games_lfp \ Repl z (Rep_game o Abs_game) = z" @@ -295,8 +299,7 @@ apply simp done -constdefs - option_of :: "(game * game) set" +definition option_of :: "(game * game) set" where "option_of \ image (\ (option, g). (Abs_game option, Abs_game g)) is_option_of" lemma option_to_is_option_of: "((option, g) \ option_of) = ((Rep_game option, Rep_game g) \ is_option_of)" @@ -437,8 +440,7 @@ qed qed -constdefs - eq_game :: "game \ game \ bool" +definition eq_game :: "game \ game \ bool" where "eq_game G H \ ge_game (G, H) \ ge_game (H, G)" lemma eq_game_sym: "(eq_game G H) = (eq_game H G)" @@ -501,9 +503,8 @@ lemma eq_game_trans: "eq_game a b \ eq_game b c \ eq_game a c" by (auto simp add: eq_game_def intro: ge_game_trans) -constdefs - zero_game :: game - "zero_game \ Game zempty zempty" +definition zero_game :: game + where "zero_game \ Game zempty zempty" consts plus_game :: "game * game \ game" @@ -838,8 +839,7 @@ then show ?thesis by blast qed -constdefs - eq_game_rel :: "(game * game) set" +definition eq_game_rel :: "(game * game) set" where "eq_game_rel \ { (p, q) . eq_game p q }" typedef Pg = "UNIV//eq_game_rel" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/ZF/HOLZF.thy --- a/src/HOL/ZF/HOLZF.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/ZF/HOLZF.thy Mon Mar 01 13:42:31 2010 +0100 @@ -19,16 +19,19 @@ Repl :: "ZF \ (ZF \ ZF) \ ZF" and Inf :: ZF -constdefs - Upair:: "ZF \ ZF \ ZF" +definition Upair :: "ZF \ ZF \ ZF" where "Upair a b == Repl (Power (Power Empty)) (% x. if x = Empty then a else b)" - Singleton:: "ZF \ ZF" + +definition Singleton:: "ZF \ ZF" where "Singleton x == Upair x x" - union :: "ZF \ ZF \ ZF" + +definition union :: "ZF \ ZF \ ZF" where "union A B == Sum (Upair A B)" - SucNat:: "ZF \ ZF" + +definition SucNat:: "ZF \ ZF" where "SucNat x == union x (Singleton x)" - subset :: "ZF \ ZF \ bool" + +definition subset :: "ZF \ ZF \ bool" where "subset A B == ! x. Elem x A \ Elem x B" axioms @@ -40,8 +43,7 @@ Regularity: "A \ Empty \ (? x. Elem x A & (! y. Elem y x \ Not (Elem y A)))" Infinity: "Elem Empty Inf & (! x. Elem x Inf \ Elem (SucNat x) Inf)" -constdefs - Sep:: "ZF \ (ZF \ bool) \ ZF" +definition Sep :: "ZF \ (ZF \ bool) \ ZF" where "Sep A p == (if (!x. Elem x A \ Not (p x)) then Empty else (let z = (\ x. Elem x A & p x) in let f = % x. (if p x then x else z) in Repl A f))" @@ -70,8 +72,7 @@ lemma Singleton: "Elem x (Singleton y) = (x = y)" by (simp add: Singleton_def Upair) -constdefs - Opair:: "ZF \ ZF \ ZF" +definition Opair :: "ZF \ ZF \ ZF" where "Opair a b == Upair (Upair a a) (Upair a b)" lemma Upair_singleton: "(Upair a a = Upair c d) = (a = c & a = d)" @@ -99,17 +100,16 @@ done qed -constdefs - Replacement :: "ZF \ (ZF \ ZF option) \ ZF" +definition Replacement :: "ZF \ (ZF \ ZF option) \ ZF" where "Replacement A f == Repl (Sep A (% a. f a \ None)) (the o f)" theorem Replacement: "Elem y (Replacement A f) = (? x. Elem x A & f x = Some y)" by (auto simp add: Replacement_def Repl Sep) -constdefs - Fst :: "ZF \ ZF" +definition Fst :: "ZF \ ZF" where "Fst q == SOME x. ? y. q = Opair x y" - Snd :: "ZF \ ZF" + +definition Snd :: "ZF \ ZF" where "Snd q == SOME y. ? x. q = Opair x y" theorem Fst: "Fst (Opair x y) = x" @@ -124,8 +124,7 @@ apply (simp_all add: Opair) done -constdefs - isOpair :: "ZF \ bool" +definition isOpair :: "ZF \ bool" where "isOpair q == ? x y. q = Opair x y" lemma isOpair: "isOpair (Opair x y) = True" @@ -134,8 +133,7 @@ lemma FstSnd: "isOpair x \ Opair (Fst x) (Snd x) = x" by (auto simp add: isOpair_def Fst Snd) -constdefs - CartProd :: "ZF \ ZF \ ZF" +definition CartProd :: "ZF \ ZF \ ZF" where "CartProd A B == Sum(Repl A (% a. Repl B (% b. Opair a b)))" lemma CartProd: "Elem x (CartProd A B) = (? a b. Elem a A & Elem b B & x = (Opair a b))" @@ -144,8 +142,7 @@ apply (auto simp add: Repl) done -constdefs - explode :: "ZF \ ZF set" +definition explode :: "ZF \ ZF set" where "explode z == { x. Elem x z }" lemma explode_Empty: "(explode x = {}) = (x = Empty)" @@ -163,10 +160,10 @@ lemma explode_Repl_eq: "explode (Repl A f) = image f (explode A)" by (simp add: explode_def Repl image_def) -constdefs - Domain :: "ZF \ ZF" +definition Domain :: "ZF \ ZF" where "Domain f == Replacement f (% p. if isOpair p then Some (Fst p) else None)" - Range :: "ZF \ ZF" + +definition Range :: "ZF \ ZF" where "Range f == Replacement f (% p. if isOpair p then Some (Snd p) else None)" theorem Domain: "Elem x (Domain f) = (? y. Elem (Opair x y) f)" @@ -188,20 +185,16 @@ theorem union: "Elem x (union A B) = (Elem x A | Elem x B)" by (auto simp add: union_def Sum Upair) -constdefs - Field :: "ZF \ ZF" +definition Field :: "ZF \ ZF" where "Field A == union (Domain A) (Range A)" -constdefs - app :: "ZF \ ZF => ZF" (infixl "\" 90) --{*function application*} +definition app :: "ZF \ ZF => ZF" (infixl "\" 90) --{*function application*} where "f \ x == (THE y. Elem (Opair x y) f)" -constdefs - isFun :: "ZF \ bool" +definition isFun :: "ZF \ bool" where "isFun f == (! x y1 y2. Elem (Opair x y1) f & Elem (Opair x y2) f \ y1 = y2)" -constdefs - Lambda :: "ZF \ (ZF \ ZF) \ ZF" +definition Lambda :: "ZF \ (ZF \ ZF) \ ZF" where "Lambda A f == Repl A (% x. Opair x (f x))" lemma Lambda_app: "Elem x A \ (Lambda A f)\x = f x" @@ -238,10 +231,10 @@ done qed -constdefs - PFun :: "ZF \ ZF \ ZF" +definition PFun :: "ZF \ ZF \ ZF" where "PFun A B == Sep (Power (CartProd A B)) isFun" - Fun :: "ZF \ ZF \ ZF" + +definition Fun :: "ZF \ ZF \ ZF" where "Fun A B == Sep (PFun A B) (\ f. Domain f = A)" lemma Fun_Range: "Elem f (Fun U V) \ subset (Range f) V" @@ -343,8 +336,7 @@ qed -constdefs - is_Elem_of :: "(ZF * ZF) set" +definition is_Elem_of :: "(ZF * ZF) set" where "is_Elem_of == { (a,b) | a b. Elem a b }" lemma cond_wf_Elem: @@ -417,8 +409,7 @@ nat2Nat_0[intro]: "nat2Nat 0 = Empty" nat2Nat_Suc[intro]: "nat2Nat (Suc n) = SucNat (nat2Nat n)" -constdefs - Nat2nat :: "ZF \ nat" +definition Nat2nat :: "ZF \ nat" where "Nat2nat == inv nat2Nat" lemma Elem_nat2Nat_inf[intro]: "Elem (nat2Nat n) Inf" @@ -426,9 +417,8 @@ apply (simp_all add: Infinity) done -constdefs - Nat :: ZF - "Nat == Sep Inf (\ N. ? n. nat2Nat n = N)" +definition Nat :: ZF + where "Nat == Sep Inf (\ N. ? n. nat2Nat n = N)" lemma Elem_nat2Nat_Nat[intro]: "Elem (nat2Nat n) Nat" by (auto simp add: Nat_def Sep) @@ -664,8 +654,7 @@ qed qed -constdefs - SpecialR :: "(ZF * ZF) set" +definition SpecialR :: "(ZF * ZF) set" where "SpecialR \ { (x, y) . x \ Empty \ y = Empty}" lemma "wf SpecialR" @@ -673,8 +662,7 @@ apply (auto simp add: SpecialR_def) done -constdefs - Ext :: "('a * 'b) set \ 'b \ 'a set" +definition Ext :: "('a * 'b) set \ 'b \ 'a set" where "Ext R y \ { x . (x, y) \ R }" lemma Ext_Elem: "Ext is_Elem_of = explode" @@ -690,8 +678,7 @@ then show "False" by (simp add: UNIV_is_not_in_ZF) qed -constdefs - implode :: "ZF set \ ZF" +definition implode :: "ZF set \ ZF" where "implode == inv explode" lemma inj_explode: "inj explode" @@ -700,12 +687,13 @@ lemma implode_explode[simp]: "implode (explode x) = x" by (simp add: implode_def inj_explode) -constdefs - regular :: "(ZF * ZF) set \ bool" +definition regular :: "(ZF * ZF) set \ bool" where "regular R == ! A. A \ Empty \ (? x. Elem x A & (! y. (y, x) \ R \ Not (Elem y A)))" - set_like :: "(ZF * ZF) set \ bool" + +definition set_like :: "(ZF * ZF) set \ bool" where "set_like R == ! y. Ext R y \ range explode" - wfzf :: "(ZF * ZF) set \ bool" + +definition wfzf :: "(ZF * ZF) set \ bool" where "wfzf R == regular R & set_like R" lemma regular_Elem: "regular is_Elem_of" @@ -717,8 +705,7 @@ lemma wfzf_is_Elem_of: "wfzf is_Elem_of" by (auto simp add: wfzf_def regular_Elem set_like_Elem) -constdefs - SeqSum :: "(nat \ ZF) \ ZF" +definition SeqSum :: "(nat \ ZF) \ ZF" where "SeqSum f == Sum (Repl Nat (f o Nat2nat))" lemma SeqSum: "Elem x (SeqSum f) = (? n. Elem x (f n))" @@ -727,8 +714,7 @@ apply auto done -constdefs - Ext_ZF :: "(ZF * ZF) set \ ZF \ ZF" +definition Ext_ZF :: "(ZF * ZF) set \ ZF \ ZF" where "Ext_ZF R s == implode (Ext R s)" lemma Elem_implode: "A \ range explode \ Elem x (implode A) = (x \ A)" @@ -750,8 +736,7 @@ "Ext_ZF_n R s 0 = Ext_ZF R s" "Ext_ZF_n R s (Suc n) = Sum (Repl (Ext_ZF_n R s n) (Ext_ZF R))" -constdefs - Ext_ZF_hull :: "(ZF * ZF) set \ ZF \ ZF" +definition Ext_ZF_hull :: "(ZF * ZF) set \ ZF \ ZF" where "Ext_ZF_hull R s == SeqSum (Ext_ZF_n R s)" lemma Elem_Ext_ZF_hull: diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/ZF/LProd.thy --- a/src/HOL/ZF/LProd.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/ZF/LProd.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/ZF/LProd.thy - ID: $Id$ Author: Steven Obua Introduces the lprod relation. @@ -95,10 +94,10 @@ show ?thesis by (auto intro: lprod) qed -constdefs - gprod_2_2 :: "('a * 'a) set \ (('a * 'a) * ('a * 'a)) set" +definition gprod_2_2 :: "('a * 'a) set \ (('a * 'a) * ('a * 'a)) set" where "gprod_2_2 R \ { ((a,b), (c,d)) . (a = c \ (b,d) \ R) \ (b = d \ (a,c) \ R) }" - gprod_2_1 :: "('a * 'a) set \ (('a * 'a) * ('a * 'a)) set" + +definition gprod_2_1 :: "('a * 'a) set \ (('a * 'a) * ('a * 'a)) set" where "gprod_2_1 R \ { ((a,b), (c,d)) . (a = d \ (b,c) \ R) \ (b = c \ (a,d) \ R) }" lemma lprod_2_3: "(a, b) \ R \ ([a, c], [b, c]) \ lprod R" @@ -170,8 +169,7 @@ apply (simp add: z' lprod_2_4) done -constdefs - perm :: "('a \ 'a) \ 'a set \ bool" +definition perm :: "('a \ 'a) \ 'a set \ bool" where "perm f A \ inj_on f A \ f ` A = A" lemma "((as,bs) \ lprod R) = @@ -183,6 +181,4 @@ lemma "trans R \ (ah@a#at, bh@b#bt) \ lprod R \ (b, a) \ R \ a = b \ (ah@at, bh@bt) \ lprod R" oops - - end \ No newline at end of file diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/ZF/MainZF.thy --- a/src/HOL/ZF/MainZF.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/ZF/MainZF.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/ZF/MainZF.thy - ID: $Id$ Author: Steven Obua Starting point for using HOLZF. @@ -9,4 +8,5 @@ theory MainZF imports Zet LProd begin + end \ No newline at end of file diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/ZF/Zet.thy --- a/src/HOL/ZF/Zet.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/ZF/Zet.thy Mon Mar 01 13:42:31 2010 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/ZF/Zet.thy - ID: $Id$ Author: Steven Obua Introduces a type 'a zet of ZF representable sets. @@ -13,15 +12,13 @@ typedef 'a zet = "{A :: 'a set | A f z. inj_on f A \ f ` A \ explode z}" by blast -constdefs - zin :: "'a \ 'a zet \ bool" +definition zin :: "'a \ 'a zet \ bool" where "zin x A == x \ (Rep_zet A)" lemma zet_ext_eq: "(A = B) = (! x. zin x A = zin x B)" by (auto simp add: Rep_zet_inject[symmetric] zin_def) -constdefs - zimage :: "('a \ 'b) \ 'a zet \ 'b zet" +definition zimage :: "('a \ 'b) \ 'a zet \ 'b zet" where "zimage f A == Abs_zet (image f (Rep_zet A))" lemma zet_def': "zet = {A :: 'a set | A f z. inj_on f A \ f ` A = explode z}" @@ -74,10 +71,10 @@ lemma zimage_iff: "zin y (zimage f A) = (? x. zin x A & y = f x)" by (auto simp add: zin_def Rep_zimage_eq) -constdefs - zimplode :: "ZF zet \ ZF" +definition zimplode :: "ZF zet \ ZF" where "zimplode A == implode (Rep_zet A)" - zexplode :: "ZF \ ZF zet" + +definition zexplode :: "ZF \ ZF zet" where "zexplode z == Abs_zet (explode z)" lemma Rep_zet_eq_explode: "? z. Rep_zet A = explode z" @@ -114,10 +111,10 @@ apply (simp_all add: comp_image_eq zet_image_mem Rep_zet) done -constdefs - zunion :: "'a zet \ 'a zet \ 'a zet" +definition zunion :: "'a zet \ 'a zet \ 'a zet" where "zunion a b \ Abs_zet ((Rep_zet a) \ (Rep_zet b))" - zsubset :: "'a zet \ 'a zet \ bool" + +definition zsubset :: "'a zet \ 'a zet \ bool" where "zsubset a b \ ! x. zin x a \ zin x b" lemma explode_union: "explode (union a b) = (explode a) \ (explode b)" @@ -181,8 +178,7 @@ apply (simp_all add: zin_def Rep_zet range_explode_eq_zet) done -constdefs - zempty :: "'a zet" +definition zempty :: "'a zet" where "zempty \ Abs_zet {}" lemma zempty[simp]: "\ (zin x zempty)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/ex/Refute_Examples.thy --- a/src/HOL/ex/Refute_Examples.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/ex/Refute_Examples.thy Mon Mar 01 13:42:31 2010 +0100 @@ -249,12 +249,13 @@ text {* "The transitive closure 'T' of an arbitrary relation 'P' is non-empty." *} -constdefs - "trans" :: "('a \ 'a \ bool) \ bool" +definition "trans" :: "('a \ 'a \ bool) \ bool" where "trans P == (ALL x y z. P x y \ P y z \ P x z)" - "subset" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" + +definition "subset" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" where "subset P Q == (ALL x y. P x y \ Q x y)" - "trans_closure" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" + +definition "trans_closure" :: "('a \ 'a \ bool) \ ('a \ 'a \ bool) \ bool" where "trans_closure P Q == (subset Q P) & (trans P) & (ALL R. subset Q R \ trans R \ subset P R)" lemma "trans_closure T P \ (\x y. T x y)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/HOL/ex/Sudoku.thy --- a/src/HOL/ex/Sudoku.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/HOL/ex/Sudoku.thy Mon Mar 01 13:42:31 2010 +0100 @@ -23,8 +23,7 @@ datatype digit = A ("1") | B ("2") | C ("3") | D ("4") | E ("5") | F ("6") | G ("7") | H ("8") | I ("9") -constdefs - valid :: "digit => digit => digit => digit => digit => digit => digit => digit => digit => bool" +definition valid :: "digit => digit => digit => digit => digit => digit => digit => digit => digit => bool" where "valid x1 x2 x3 x4 x5 x6 x7 x8 x9 == (x1 \ x2) \ (x1 \ x3) \ (x1 \ x4) \ (x1 \ x5) \ (x1 \ x6) \ (x1 \ x7) \ (x1 \ x8) \ (x1 \ x9) @@ -36,8 +35,7 @@ \ (x7 \ x8) \ (x7 \ x9) \ (x8 \ x9)" -constdefs - sudoku :: "digit => digit => digit => digit => digit => digit => digit => digit => digit => +definition sudoku :: "digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => @@ -45,7 +43,7 @@ digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => digit => - digit => digit => digit => digit => digit => digit => digit => digit => digit => bool" + digit => digit => digit => digit => digit => digit => digit => digit => digit => bool" where "sudoku x11 x12 x13 x14 x15 x16 x17 x18 x19 x21 x22 x23 x24 x25 x26 x27 x28 x29 diff -r 1810b1ade437 -r 47ee18b6ae32 src/Sequents/LK0.thy --- a/src/Sequents/LK0.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/Sequents/LK0.thy Mon Mar 01 13:42:31 2010 +0100 @@ -122,8 +122,7 @@ The: "[| $H |- $E, P(a), $F; !!x.$H, P(x) |- $E, x=a, $F |] ==> $H |- $E, P(THE x. P(x)), $F" -constdefs - If :: "[o, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10) +definition If :: "[o, 'a, 'a] => 'a" ("(if (_)/ then (_)/ else (_))" 10) where "If(P,x,y) == THE z::'a. (P --> z=x) & (~P --> z=y)" diff -r 1810b1ade437 -r 47ee18b6ae32 src/ZF/Sum.thy --- a/src/ZF/Sum.thy Mon Mar 01 12:30:55 2010 +0100 +++ b/src/ZF/Sum.thy Mon Mar 01 13:42:31 2010 +0100 @@ -11,21 +11,20 @@ global -constdefs - sum :: "[i,i]=>i" (infixr "+" 65) +definition sum :: "[i,i]=>i" (infixr "+" 65) where "A+B == {0}*A Un {1}*B" - Inl :: "i=>i" +definition Inl :: "i=>i" where "Inl(a) == <0,a>" - Inr :: "i=>i" +definition Inr :: "i=>i" where "Inr(b) == <1,b>" - "case" :: "[i=>i, i=>i, i]=>i" +definition "case" :: "[i=>i, i=>i, i]=>i" where "case(c,d) == (%. cond(y, d(z), c(z)))" (*operator for selecting out the various summands*) - Part :: "[i,i=>i] => i" +definition Part :: "[i,i=>i] => i" where "Part(A,h) == {x: A. EX z. x = h(z)}" local