replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
--- a/doc-src/TutorialI/Advanced/Partial.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/TutorialI/Advanced/Partial.thy Mon Mar 01 13:40:23 2010 +0100
@@ -34,7 +34,7 @@
preconditions:
*}
-constdefs subtract :: "nat \<Rightarrow> nat \<Rightarrow> nat"
+definition subtract :: "nat \<Rightarrow> nat \<Rightarrow> nat" where
"n \<le> m \<Longrightarrow> subtract m n \<equiv> m - n"
text{*
@@ -85,7 +85,7 @@
Phrased differently, the relation
*}
-constdefs step1 :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<times> 'a)set"
+definition step1 :: "('a \<Rightarrow> 'a) \<Rightarrow> ('a \<times> 'a)set" where
"step1 f \<equiv> {(y,x). y = f x \<and> y \<noteq> x}"
text{*\noindent
@@ -160,7 +160,7 @@
consider the following definition of function @{const find}:
*}
-constdefs find2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
+definition find2 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" where
"find2 f x \<equiv>
fst(while (\<lambda>(x,x'). x' \<noteq> x) (\<lambda>(x,x'). (x',f x')) (x,f x))"
--- a/doc-src/TutorialI/CTL/CTL.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/TutorialI/CTL/CTL.thy Mon Mar 01 13:40:23 2010 +0100
@@ -365,8 +365,7 @@
*}
(*<*)
-constdefs
- eufix :: "state set \<Rightarrow> state set \<Rightarrow> state set \<Rightarrow> state set"
+definition eufix :: "state set \<Rightarrow> state set \<Rightarrow> state set \<Rightarrow> state set" where
"eufix A B T \<equiv> B \<union> A \<inter> (M\<inverse> `` T)"
lemma "lfp(eufix A B) \<subseteq> eusem A B"
@@ -397,8 +396,7 @@
done
(*
-constdefs
- eusem :: "state set \<Rightarrow> state set \<Rightarrow> state set"
+definition eusem :: "state set \<Rightarrow> state set \<Rightarrow> state set" where
"eusem A B \<equiv> {s. \<exists>p\<in>Paths s. \<exists>j. p j \<in> B \<and> (\<forall>i < j. p i \<in> A)}"
axioms
@@ -414,8 +412,7 @@
apply(blast intro: M_total[THEN someI_ex])
done
-constdefs
- pcons :: "state \<Rightarrow> (nat \<Rightarrow> state) \<Rightarrow> (nat \<Rightarrow> state)"
+definition pcons :: "state \<Rightarrow> (nat \<Rightarrow> state) \<Rightarrow> (nat \<Rightarrow> state)" where
"pcons s p == \<lambda>i. case i of 0 \<Rightarrow> s | Suc j \<Rightarrow> p j"
lemma pcons_PathI: "[| (s,t) : M; p \<in> Paths t |] ==> pcons s p \<in> Paths s";
--- a/doc-src/TutorialI/Misc/Option2.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/TutorialI/Misc/Option2.thy Mon Mar 01 13:40:23 2010 +0100
@@ -24,8 +24,7 @@
*}
(*<*)
(*
-constdefs
- infplus :: "nat option \<Rightarrow> nat option \<Rightarrow> nat option"
+definition infplus :: "nat option \<Rightarrow> nat option \<Rightarrow> nat option" where
"infplus x y \<equiv> case x of None \<Rightarrow> None
| Some m \<Rightarrow> (case y of None \<Rightarrow> None | Some n \<Rightarrow> Some(m+n))"
--- a/doc-src/TutorialI/Overview/LNCS/FP1.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/TutorialI/Overview/LNCS/FP1.thy Mon Mar 01 13:40:23 2010 +0100
@@ -62,7 +62,7 @@
consts xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"
defs xor_def: "xor x y \<equiv> x \<and> \<not>y \<or> \<not>x \<and> y"
-constdefs nand :: "bool \<Rightarrow> bool \<Rightarrow> bool"
+definition nand :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
"nand x y \<equiv> \<not>(x \<and> y)"
lemma "\<not> xor x x"
--- a/doc-src/TutorialI/Overview/LNCS/Ordinal.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/TutorialI/Overview/LNCS/Ordinal.thy Mon Mar 01 13:40:23 2010 +0100
@@ -9,8 +9,7 @@
"pred (Succ a) n = Some a"
"pred (Limit f) n = Some (f n)"
-constdefs
- OpLim :: "(nat \<Rightarrow> (ordinal \<Rightarrow> ordinal)) \<Rightarrow> (ordinal \<Rightarrow> ordinal)"
+definition OpLim :: "(nat \<Rightarrow> (ordinal \<Rightarrow> ordinal)) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" where
"OpLim F a \<equiv> Limit (\<lambda>n. F n a)"
OpItw :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" ("\<Squnion>")
"\<Squnion>f \<equiv> OpLim (power f)"
@@ -29,8 +28,7 @@
"\<nabla>f (Succ a) = f (Succ (\<nabla>f a))"
"\<nabla>f (Limit h) = Limit (\<lambda>n. \<nabla>f (h n))"
-constdefs
- deriv :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)"
+definition deriv :: "(ordinal \<Rightarrow> ordinal) \<Rightarrow> (ordinal \<Rightarrow> ordinal)" where
"deriv f \<equiv> \<nabla>(\<Squnion>f)"
consts
@@ -40,8 +38,7 @@
"veblen (Succ a) = \<nabla>(OpLim (power (veblen a)))"
"veblen (Limit f) = \<nabla>(OpLim (\<lambda>n. veblen (f n)))"
-constdefs
- veb :: "ordinal \<Rightarrow> ordinal"
+definition veb :: "ordinal \<Rightarrow> ordinal" where
"veb a \<equiv> veblen a Zero"
epsilon0 :: ordinal ("\<epsilon>\<^sub>0")
"\<epsilon>\<^sub>0 \<equiv> veb Zero"
--- a/doc-src/TutorialI/Protocol/Message.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/TutorialI/Protocol/Message.thy Mon Mar 01 13:40:23 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. \<exists>X. Crypt K X \<in> H}"
--- a/doc-src/TutorialI/Rules/Primes.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/TutorialI/Rules/Primes.thy Mon Mar 01 13:40:23 2010 +0100
@@ -99,8 +99,7 @@
(**** The material below was omitted from the book ****)
-constdefs
- is_gcd :: "[nat,nat,nat] \<Rightarrow> bool" (*gcd as a relation*)
+definition is_gcd :: "[nat,nat,nat] \<Rightarrow> bool" where (*gcd as a relation*)
"is_gcd p m n == p dvd m \<and> p dvd n \<and>
(ALL d. d dvd m \<and> d dvd n \<longrightarrow> d dvd p)"
--- a/doc-src/TutorialI/Sets/Examples.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/TutorialI/Sets/Examples.thy Mon Mar 01 13:40:23 2010 +0100
@@ -156,8 +156,7 @@
lemma "{x. P x \<longrightarrow> Q x} = -{x. P x} \<union> {x. Q x}"
by blast
-constdefs
- prime :: "nat set"
+definition prime :: "nat set" where
"prime == {p. 1<p & (ALL m. m dvd p --> m=1 | m=p)}"
lemma "{p*q | p q. p\<in>prime \<and> q\<in>prime} =
--- a/doc-src/ZF/If.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/ZF/If.thy Mon Mar 01 13:40:23 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:
--- a/doc-src/ZF/ZF_examples.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/doc-src/ZF/ZF_examples.thy Mon Mar 01 13:40:23 2010 +0100
@@ -64,7 +64,7 @@
"t \<in> bt(A) ==> \<forall>k \<in> 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 \<in> bt(A) ==> n_nodes_tail(t) = n_nodes(t)"
--- a/src/FOL/IFOL.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/FOL/IFOL.thy Mon Mar 01 13:40:23 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
--- a/src/FOL/ex/If.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/FOL/ex/If.thy Mon Mar 01 13:40:23 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:
--- a/src/FOLP/ex/If.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/FOLP/ex/If.thy Mon Mar 01 13:40:23 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:
--- a/src/HOL/Algebra/AbelCoset.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Algebra/AbelCoset.thy Mon Mar 01 13:40:23 2010 +0100
@@ -38,15 +38,12 @@
("racong\<index> _")
"a_r_congruent G \<equiv> r_congruent \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>"
-constdefs
- A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) monoid"
- (infixl "A'_Mod" 65)
+definition A_FactGroup :: "[('a,'b) ring_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "A'_Mod" 65) where
--{*Actually defined for groups rather than monoids*}
"A_FactGroup G H \<equiv> FactGroup \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr> H"
-constdefs
- a_kernel :: "('a, 'm) ring_scheme \<Rightarrow> ('b, 'n) ring_scheme \<Rightarrow>
- ('a \<Rightarrow> 'b) \<Rightarrow> 'a set"
+definition a_kernel :: "('a, 'm) ring_scheme \<Rightarrow> ('b, 'n) ring_scheme \<Rightarrow>
+ ('a \<Rightarrow> 'b) \<Rightarrow> 'a set" where
--{*the kernel of a homomorphism (additive)*}
"a_kernel G H h \<equiv> kernel \<lparr>carrier = carrier G, mult = add G, one = zero G\<rparr>
\<lparr>carrier = carrier H, mult = add H, one = zero H\<rparr> h"
--- a/src/HOL/Algebra/Bij.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Algebra/Bij.thy Mon Mar 01 13:40:23 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 \<Rightarrow> ('a \<Rightarrow> 'a) set"
+definition Bij :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) set" where
--{*Only extensional functions, since otherwise we get too many.*}
"Bij S \<equiv> extensional S \<inter> {f. bij_betw f S S}"
- BijGroup :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
+definition BijGroup :: "'a set \<Rightarrow> ('a \<Rightarrow> 'a) monoid" where
"BijGroup S \<equiv>
\<lparr>carrier = Bij S,
mult = \<lambda>g \<in> Bij S. \<lambda>f \<in> Bij S. compose S g f,
@@ -71,11 +69,10 @@
done
-constdefs
- auto :: "('a, 'b) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) set"
+definition auto :: "('a, 'b) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) set" where
"auto G \<equiv> hom G G \<inter> Bij (carrier G)"
- AutoGroup :: "('a, 'c) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) monoid"
+definition AutoGroup :: "('a, 'c) monoid_scheme \<Rightarrow> ('a \<Rightarrow> 'a) monoid" where
"AutoGroup G \<equiv> BijGroup (carrier G) \<lparr>carrier := auto G\<rparr>"
lemma (in group) id_in_auto: "(\<lambda>x \<in> carrier G. x) \<in> auto G"
--- a/src/HOL/Algebra/Coset.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Algebra/Coset.thy Mon Mar 01 13:40:23 2010 +0100
@@ -751,8 +751,7 @@
subsection {*Order of a Group and Lagrange's Theorem*}
-constdefs
- order :: "('a, 'b) monoid_scheme \<Rightarrow> nat"
+definition order :: "('a, 'b) monoid_scheme \<Rightarrow> nat" where
"order S \<equiv> 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] \<Rightarrow> ('a set) monoid"
- (infixl "Mod" 65)
+definition FactGroup :: "[('a,'b) monoid_scheme, 'a set] \<Rightarrow> ('a set) monoid" (infixl "Mod" 65) where
--{*Actually defined for groups rather than monoids*}
"FactGroup G H \<equiv>
\<lparr>carrier = rcosets\<^bsub>G\<^esub> H, mult = set_mult G, one = H\<rparr>"
@@ -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 \<Rightarrow> ('b, 'n) monoid_scheme \<Rightarrow>
- ('a \<Rightarrow> 'b) \<Rightarrow> 'a set"
+definition kernel :: "('a, 'm) monoid_scheme \<Rightarrow> ('b, 'n) monoid_scheme \<Rightarrow>
+ ('a \<Rightarrow> 'b) \<Rightarrow> 'a set" where
--{*the kernel of a homomorphism*}
"kernel G H h \<equiv> {x. x \<in> carrier G & h x = \<one>\<^bsub>H\<^esub>}"
--- a/src/HOL/Algebra/Divisibility.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Algebra/Divisibility.thy Mon Mar 01 13:40:23 2010 +0100
@@ -3630,8 +3630,7 @@
text {* Number of factors for wellfoundedness *}
-constdefs
- factorcount :: "_ \<Rightarrow> 'a \<Rightarrow> nat"
+definition factorcount :: "_ \<Rightarrow> 'a \<Rightarrow> nat" where
"factorcount G a == THE c. (ALL as. set as \<subseteq> carrier G \<and>
wfactors G as a \<longrightarrow> c = length as)"
--- a/src/HOL/Algebra/FiniteProduct.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Algebra/FiniteProduct.thy Mon Mar 01 13:40:23 2010 +0100
@@ -26,8 +26,7 @@
inductive_cases empty_foldSetDE [elim!]: "({}, x) \<in> 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) \<in> foldSetD D f e"
lemma foldSetD_closed:
--- a/src/HOL/Algebra/Group.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Algebra/Group.thy Mon Mar 01 13:40:23 2010 +0100
@@ -478,8 +478,7 @@
subsection {* Direct Products *}
-constdefs
- DirProd :: "_ \<Rightarrow> _ \<Rightarrow> ('a \<times> 'b) monoid" (infixr "\<times>\<times>" 80)
+definition DirProd :: "_ \<Rightarrow> _ \<Rightarrow> ('a \<times> 'b) monoid" (infixr "\<times>\<times>" 80) where
"G \<times>\<times> H \<equiv> \<lparr>carrier = carrier G \<times> carrier H,
mult = (\<lambda>(g, h) (g', h'). (g \<otimes>\<^bsub>G\<^esub> g', h \<otimes>\<^bsub>H\<^esub> h')),
one = (\<one>\<^bsub>G\<^esub>, \<one>\<^bsub>H\<^esub>)\<rparr>"
@@ -545,8 +544,7 @@
"[|h \<in> hom G H; i \<in> hom H I|] ==> compose (carrier G) i h \<in> hom G I"
by (fastsimp simp add: hom_def compose_def)
-constdefs
- iso :: "_ => _ => ('a => 'b) set" (infixr "\<cong>" 60)
+definition iso :: "_ => _ => ('a => 'b) set" (infixr "\<cong>" 60) where
"G \<cong> H == {h. h \<in> hom G H & bij_betw h (carrier G) (carrier H)}"
lemma iso_refl: "(%x. x) \<in> G \<cong> G"
--- a/src/HOL/Algebra/IntRing.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Algebra/IntRing.thy Mon Mar 01 13:40:23 2010 +0100
@@ -22,8 +22,7 @@
subsection {* @{text "\<Z>"}: The Set of Integers as Algebraic Structure *}
-constdefs
- int_ring :: "int ring" ("\<Z>")
+definition int_ring :: "int ring" ("\<Z>") where
"int_ring \<equiv> \<lparr>carrier = UNIV, mult = op *, one = 1, zero = 0, add = op +\<rparr>"
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>\<Z>\<^esub> {k}) +>\<^bsub>\<Z>\<^esub> r"
lemmas ZMod_defs =
@@ -407,8 +405,7 @@
subsection {* Factorization *}
-constdefs
- ZFact :: "int \<Rightarrow> int set ring"
+definition ZFact :: "int \<Rightarrow> int set ring" where
"ZFact k == \<Z> Quot (Idl\<^bsub>\<Z>\<^esub> {k})"
lemmas ZFact_defs = ZFact_def FactRing_def
--- a/src/HOL/Algebra/Ring.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Algebra/Ring.thy Mon Mar 01 13:40:23 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"
--- a/src/HOL/Auth/CertifiedEmail.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/CertifiedEmail.thy Mon Mar 01 13:40:23 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|}"
--- a/src/HOL/Auth/Guard/Extensions.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Guard/Extensions.thy Mon Mar 01 13:40:23 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"
--- a/src/HOL/Auth/Guard/Guard.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Guard/Guard.thy Mon Mar 01 13:40:23 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]
--- a/src/HOL/Auth/Guard/GuardK.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Guard/GuardK.thy Mon Mar 01 13:40:23 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]
--- a/src/HOL/Auth/Guard/Guard_Public.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Guard/Guard_Public.thy Mon Mar 01 13:40:23 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 |] ==>
--- a/src/HOL/Auth/Guard/Guard_Shared.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Guard/Guard_Shared.thy Mon Mar 01 13:40:23 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 |] ==>
--- a/src/HOL/Auth/Guard/Guard_Yahalom.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy Mon Mar 01 13:40:23 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 |] ==>
--- a/src/HOL/Auth/Guard/P1.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Guard/P1.thy Mon Mar 01 13:40:23 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
--- a/src/HOL/Auth/Guard/P2.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Guard/P2.thy Mon Mar 01 13:40:23 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'
--- a/src/HOL/Auth/Guard/Proto.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Guard/Proto.thy Mon Mar 01 13:40:23 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"
--- a/src/HOL/Auth/KerberosIV.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/KerberosIV.thy Mon Mar 01 13:40:23 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 ==
\<exists>A B Ts.
Says Tgs A (Crypt authK
--- a/src/HOL/Auth/KerberosIV_Gets.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/KerberosIV_Gets.thy Mon Mar 01 13:40:23 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 ==
\<exists>A B Ts.
Says Tgs A (Crypt authK
--- a/src/HOL/Auth/KerberosV.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/KerberosV.thy Mon Mar 01 13:40:23 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 ==
\<exists>A B tt.
Says Tgs A \<lbrace>Crypt authK \<lbrace>Key servK, Agent B, tt\<rbrace>,
--- a/src/HOL/Auth/Message.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Message.thy Mon Mar 01 13:40:23 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. \<exists>X. Crypt K X \<in> H}"
--- a/src/HOL/Auth/Smartcard/ShoupRubin.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Smartcard/ShoupRubin.thy Mon Mar 01 13:40:23 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 \<in> sr \<Longrightarrow> secureM"
-constdefs
-
- Unique :: "[event, event list] => bool" ("Unique _ on _")
+definition Unique :: "[event, event list] => bool" ("Unique _ on _") where
"Unique ev on evs ==
ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
--- a/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Smartcard/ShoupRubinBella.thy Mon Mar 01 13:40:23 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 \<in> srb \<Longrightarrow> secureM"
-constdefs
-
- Unique :: "[event, event list] => bool" ("Unique _ on _")
+definition Unique :: "[event, event list] => bool" ("Unique _ on _") where
"Unique ev on evs ==
ev \<notin> set (tl (dropWhile (% z. z \<noteq> ev) evs))"
--- a/src/HOL/Auth/Smartcard/Smartcard.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Smartcard/Smartcard.thy Mon Mar 01 13:40:23 2010 +0100
@@ -43,15 +43,11 @@
shrK_disj_pin [iff]: "shrK P \<noteq> pin Q"
crdK_disj_pin [iff]: "crdK C \<noteq> pin P"
-constdefs
- legalUse :: "card => bool" ("legalUse (_)")
+definition legalUse :: "card => bool" ("legalUse (_)") where
"legalUse C == C \<notin> stolen"
-consts
- illegalUse :: "card => bool"
-primrec
- illegalUse_def:
- "illegalUse (Card A) = ( (Card A \<in> stolen \<and> A \<in> bad) \<or> Card A \<in> cloned )"
+primrec illegalUse :: "card => bool" where
+ illegalUse_def: "illegalUse (Card A) = ( (Card A \<in> stolen \<and> A \<in> bad) \<or> Card A \<in> cloned )"
text{*initState must be defined with care*}
--- a/src/HOL/Auth/TLS.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/TLS.thy Mon Mar 01 13:40:23 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
--- a/src/HOL/Auth/Yahalom.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/Yahalom.thy Mon Mar 01 13:40:23 2010 +0100
@@ -74,8 +74,7 @@
==> Notes Spy {|Nonce NA, Nonce NB, Key K|} # evso \<in> yahalom"
-constdefs
- KeyWithNonce :: "[key, nat, event list] => bool"
+definition KeyWithNonce :: "[key, nat, event list] => bool" where
"KeyWithNonce K NB evs ==
\<exists>A B na X.
Says Server A {|Crypt (shrK A) {|Agent B, Key K, na, Nonce NB|}, X|}
--- a/src/HOL/Auth/ZhouGollmann.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Auth/ZhouGollmann.thy Mon Mar 01 13:40:23 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}"
--- a/src/HOL/Bali/AxCompl.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/AxCompl.thy Mon Mar 01 13:40:23 2010 +0100
@@ -20,9 +20,7 @@
section "set of not yet initialzed classes"
-constdefs
-
- nyinitcls :: "prog \<Rightarrow> state \<Rightarrow> qtname set"
+definition nyinitcls :: "prog \<Rightarrow> state \<Rightarrow> qtname set" where
"nyinitcls G s \<equiv> {C. is_class G C \<and> \<not> initd C s}"
lemma nyinitcls_subset_class: "nyinitcls G s \<subseteq> {C. is_class G C}"
@@ -115,8 +113,7 @@
section "init-le"
-constdefs
- init_le :: "prog \<Rightarrow> nat \<Rightarrow> state \<Rightarrow> bool" ("_\<turnstile>init\<le>_" [51,51] 50)
+definition init_le :: "prog \<Rightarrow> nat \<Rightarrow> state \<Rightarrow> bool" ("_\<turnstile>init\<le>_" [51,51] 50) where
"G\<turnstile>init\<le>n \<equiv> \<lambda>s. card (nyinitcls G s) \<le> n"
lemma init_le_def2 [simp]: "(G\<turnstile>init\<le>n) s = (card (nyinitcls G s)\<le>n)"
@@ -135,9 +132,7 @@
section "Most General Triples and Formulas"
-constdefs
-
- remember_init_state :: "state assn" ("\<doteq>")
+definition remember_init_state :: "state assn" ("\<doteq>") where
"\<doteq> \<equiv> \<lambda>Y s Z. s = Z"
lemma remember_init_state_def2 [simp]: "\<doteq> Y = op ="
@@ -579,8 +574,7 @@
unroll the loop in iterated evaluations of the expression and evaluations of
the loop body. *}
-constdefs
- unroll:: "prog \<Rightarrow> label \<Rightarrow> expr \<Rightarrow> stmt \<Rightarrow> (state \<times> state) set"
+definition unroll :: "prog \<Rightarrow> label \<Rightarrow> expr \<Rightarrow> stmt \<Rightarrow> (state \<times> state) set" where
"unroll G l e c \<equiv> {(s,t). \<exists> v s1 s2.
G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s1 \<and> the_Bool v \<and> normal s1 \<and>
--- a/src/HOL/Bali/AxExample.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/AxExample.thy Mon Mar 01 13:40:23 2010 +0100
@@ -8,8 +8,7 @@
imports AxSem Example
begin
-constdefs
- arr_inv :: "st \<Rightarrow> bool"
+definition arr_inv :: "st \<Rightarrow> bool" where
"arr_inv \<equiv> \<lambda>s. \<exists>obj a T el. globs s (Stat Base) = Some obj \<and>
values obj (Inl (arr, Base)) = Some (Addr a) \<and>
heap s a = Some \<lparr>tag=Arr T 2,values=el\<rparr>"
--- a/src/HOL/Bali/AxSem.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/AxSem.thy Mon Mar 01 13:40:23 2010 +0100
@@ -63,8 +63,7 @@
"res" <= (type) "AxSem.res"
"a assn" <= (type) "vals \<Rightarrow> state \<Rightarrow> a \<Rightarrow> bool"
-constdefs
- assn_imp :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> bool" (infixr "\<Rightarrow>" 25)
+definition assn_imp :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> bool" (infixr "\<Rightarrow>" 25) where
"P \<Rightarrow> Q \<equiv> \<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z"
lemma assn_imp_def2 [iff]: "(P \<Rightarrow> Q) = (\<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z)"
@@ -77,8 +76,7 @@
subsection "peek-and"
-constdefs
- peek_and :: "'a assn \<Rightarrow> (state \<Rightarrow> bool) \<Rightarrow> 'a assn" (infixl "\<and>." 13)
+definition peek_and :: "'a assn \<Rightarrow> (state \<Rightarrow> bool) \<Rightarrow> 'a assn" (infixl "\<and>." 13) where
"P \<and>. p \<equiv> \<lambda>Y s Z. P Y s Z \<and> p s"
lemma peek_and_def2 [simp]: "peek_and P p Y s = (\<lambda>Z. (P Y s Z \<and> p s))"
@@ -117,8 +115,7 @@
subsection "assn-supd"
-constdefs
- assn_supd :: "'a assn \<Rightarrow> (state \<Rightarrow> state) \<Rightarrow> 'a assn" (infixl ";." 13)
+definition assn_supd :: "'a assn \<Rightarrow> (state \<Rightarrow> state) \<Rightarrow> 'a assn" (infixl ";." 13) where
"P ;. f \<equiv> \<lambda>Y s' Z. \<exists>s. P Y s Z \<and> s' = f s"
lemma assn_supd_def2 [simp]: "assn_supd P f Y s' Z = (\<exists>s. P Y s Z \<and> s' = f s)"
@@ -128,8 +125,7 @@
subsection "supd-assn"
-constdefs
- supd_assn :: "(state \<Rightarrow> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr ".;" 13)
+definition supd_assn :: "(state \<Rightarrow> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr ".;" 13) where
"f .; P \<equiv> \<lambda>Y s. P Y (f s)"
@@ -148,8 +144,7 @@
subsection "subst-res"
-constdefs
- subst_res :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" ("_\<leftarrow>_" [60,61] 60)
+definition subst_res :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" ("_\<leftarrow>_" [60,61] 60) where
"P\<leftarrow>w \<equiv> \<lambda>Y. P w"
lemma subst_res_def2 [simp]: "(P\<leftarrow>w) Y = P w"
@@ -184,8 +179,7 @@
subsection "subst-Bool"
-constdefs
- subst_Bool :: "'a assn \<Rightarrow> bool \<Rightarrow> 'a assn" ("_\<leftarrow>=_" [60,61] 60)
+definition subst_Bool :: "'a assn \<Rightarrow> bool \<Rightarrow> 'a assn" ("_\<leftarrow>=_" [60,61] 60) where
"P\<leftarrow>=b \<equiv> \<lambda>Y s Z. \<exists>v. P (Val v) s Z \<and> (normal s \<longrightarrow> the_Bool v=b)"
lemma subst_Bool_def2 [simp]:
@@ -200,8 +194,7 @@
subsection "peek-res"
-constdefs
- peek_res :: "(res \<Rightarrow> 'a assn) \<Rightarrow> 'a assn"
+definition peek_res :: "(res \<Rightarrow> 'a assn) \<Rightarrow> 'a assn" where
"peek_res Pf \<equiv> \<lambda>Y. Pf Y Y"
syntax
@@ -229,8 +222,7 @@
subsection "ign-res"
-constdefs
- ign_res :: " 'a assn \<Rightarrow> 'a assn" ("_\<down>" [1000] 1000)
+definition ign_res :: " 'a assn \<Rightarrow> 'a assn" ("_\<down>" [1000] 1000) where
"P\<down> \<equiv> \<lambda>Y s Z. \<exists>Y. P Y s Z"
lemma ign_res_def2 [simp]: "P\<down> Y s Z = (\<exists>Y. P Y s Z)"
@@ -261,8 +253,7 @@
subsection "peek-st"
-constdefs
- peek_st :: "(st \<Rightarrow> 'a assn) \<Rightarrow> 'a assn"
+definition peek_st :: "(st \<Rightarrow> 'a assn) \<Rightarrow> 'a assn" where
"peek_st P \<equiv> \<lambda>Y s. P (store s) Y s"
syntax
@@ -306,8 +297,7 @@
subsection "ign-res-eq"
-constdefs
- ign_res_eq :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" ("_\<down>=_" [60,61] 60)
+definition ign_res_eq :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" ("_\<down>=_" [60,61] 60) where
"P\<down>=w \<equiv> \<lambda>Y:. P\<down> \<and>. (\<lambda>s. Y=w)"
lemma ign_res_eq_def2 [simp]: "(P\<down>=w) Y s Z = ((\<exists>Y. P Y s Z) \<and> Y=w)"
@@ -337,8 +327,7 @@
subsection "RefVar"
-constdefs
- RefVar :: "(state \<Rightarrow> vvar \<times> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn"(infixr "..;" 13)
+definition RefVar :: "(state \<Rightarrow> vvar \<times> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr "..;" 13) where
"vf ..; P \<equiv> \<lambda>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 \<Rightarrow> obj_tag \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
+definition Alloc :: "prog \<Rightarrow> obj_tag \<Rightarrow> 'a assn \<Rightarrow> 'a assn" where
"Alloc G otag P \<equiv> \<lambda>Y s Z.
\<forall>s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a\<rightarrow> s'\<longrightarrow> P (Val (Addr a)) s' Z"
- SXAlloc :: "prog \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
+definition SXAlloc :: "prog \<Rightarrow> 'a assn \<Rightarrow> 'a assn" where
"SXAlloc G P \<equiv> \<lambda>Y s Z. \<forall>s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<longrightarrow> P Y s' Z"
@@ -372,8 +360,7 @@
section "validity"
-constdefs
- type_ok :: "prog \<Rightarrow> term \<Rightarrow> state \<Rightarrow> bool"
+definition type_ok :: "prog \<Rightarrow> term \<Rightarrow> state \<Rightarrow> bool" where
"type_ok G t s \<equiv>
\<exists>L T C A. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<and>
\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A )
@@ -419,10 +406,8 @@
apply auto
done
-constdefs
- mtriples :: "('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<Rightarrow> 'sig \<Rightarrow> expr) \<Rightarrow>
- ('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<times> 'sig) set \<Rightarrow> 'a triples"
- ("{{(1_)}/ _-\<succ>/ {(1_)} | _}"[3,65,3,65]75)
+definition mtriples :: "('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<Rightarrow> 'sig \<Rightarrow> expr) \<Rightarrow>
+ ('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<times> 'sig) set \<Rightarrow> 'a triples" ("{{(1_)}/ _-\<succ>/ {(1_)} | _}"[3,65,3,65]75) where
"{{P} tf-\<succ> {Q} | ms} \<equiv> (\<lambda>(C,sig). {Normal(P C sig)} tf C sig-\<succ> {Q C sig})`ms"
consts
@@ -641,8 +626,7 @@
axioms
*)
-constdefs
- adapt_pre :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn"
+definition adapt_pre :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn" where
"adapt_pre P Q Q'\<equiv>\<lambda>Y s Z. \<forall>Y' s'. \<exists>Z'. P Y s Z' \<and> (Q Y' s' Z' \<longrightarrow> Q' Y' s' Z)"
--- a/src/HOL/Bali/Basis.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/Basis.thy Mon Mar 01 13:40:23 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 \<times> 'b) list \<Rightarrow> bool"
+definition unique :: "('a \<times> 'b) list \<Rightarrow> bool" where
"unique \<equiv> distinct \<circ> map fst"
lemma uniqueD [rule_format (no_asm)]:
--- a/src/HOL/Bali/Conform.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/Conform.thy Mon Mar 01 13:40:23 2010 +0100
@@ -22,9 +22,7 @@
section "extension of global store"
-constdefs
-
- gext :: "st \<Rightarrow> st \<Rightarrow> bool" ("_\<le>|_" [71,71] 70)
+definition gext :: "st \<Rightarrow> st \<Rightarrow> bool" ("_\<le>|_" [71,71] 70) where
"s\<le>|s' \<equiv> \<forall>r. \<forall>obj\<in>globs s r: \<exists>obj'\<in>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 \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>_" [71,71,71,71] 70)
+definition conf :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>_" [71,71,71,71] 70) where
"G,s\<turnstile>v\<Colon>\<preceq>T \<equiv> \<exists>T'\<in>typeof (\<lambda>a. Option.map obj_ty (heap s a)) v:G\<turnstile>T'\<preceq>T"
lemma conf_cong [simp]: "G,set_locals l s\<turnstile>v\<Colon>\<preceq>T = G,s\<turnstile>v\<Colon>\<preceq>T"
@@ -181,10 +177,7 @@
section "value list conformance"
-constdefs
-
- lconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool"
- ("_,_\<turnstile>_[\<Colon>\<preceq>]_" [71,71,71,71] 70)
+definition lconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool" ("_,_\<turnstile>_[\<Colon>\<preceq>]_" [71,71,71,71] 70) where
"G,s\<turnstile>vs[\<Colon>\<preceq>]Ts \<equiv> \<forall>n. \<forall>T\<in>Ts n: \<exists>v\<in>vs n: G,s\<turnstile>v\<Colon>\<preceq>T"
lemma lconfD: "\<lbrakk>G,s\<turnstile>vs[\<Colon>\<preceq>]Ts; Ts n = Some T\<rbrakk> \<Longrightarrow> G,s\<turnstile>(the (vs n))\<Colon>\<preceq>T"
@@ -267,10 +260,7 @@
*}
-constdefs
-
- wlconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool"
- ("_,_\<turnstile>_[\<sim>\<Colon>\<preceq>]_" [71,71,71,71] 70)
+definition wlconf :: "prog \<Rightarrow> st \<Rightarrow> ('a, val) table \<Rightarrow> ('a, ty) table \<Rightarrow> bool" ("_,_\<turnstile>_[\<sim>\<Colon>\<preceq>]_" [71,71,71,71] 70) where
"G,s\<turnstile>vs[\<sim>\<Colon>\<preceq>]Ts \<equiv> \<forall>n. \<forall>T\<in>Ts n: \<forall> v\<in>vs n: G,s\<turnstile>v\<Colon>\<preceq>T"
lemma wlconfD: "\<lbrakk>G,s\<turnstile>vs[\<sim>\<Colon>\<preceq>]Ts; Ts n = Some T; vs n = Some v\<rbrakk> \<Longrightarrow> G,s\<turnstile>v\<Colon>\<preceq>T"
@@ -348,9 +338,7 @@
section "object conformance"
-constdefs
-
- oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>\<surd>_" [71,71,71,71] 70)
+definition oconf :: "prog \<Rightarrow> st \<Rightarrow> obj \<Rightarrow> oref \<Rightarrow> bool" ("_,_\<turnstile>_\<Colon>\<preceq>\<surd>_" [71,71,71,71] 70) where
"G,s\<turnstile>obj\<Colon>\<preceq>\<surd>r \<equiv> G,s\<turnstile>values obj[\<Colon>\<preceq>]var_tys G (tag obj) r \<and>
(case r of
Heap a \<Rightarrow> is_type G (obj_ty obj)
@@ -385,9 +373,7 @@
section "state conformance"
-constdefs
-
- conforms :: "state \<Rightarrow> env' \<Rightarrow> bool" ( "_\<Colon>\<preceq>_" [71,71] 70)
+definition conforms :: "state \<Rightarrow> env' \<Rightarrow> bool" ("_\<Colon>\<preceq>_" [71,71] 70) where
"xs\<Colon>\<preceq>E \<equiv> let (G, L) = E; s = snd xs; l = locals s in
(\<forall>r. \<forall>obj\<in>globs s r: G,s\<turnstile>obj \<Colon>\<preceq>\<surd>r) \<and>
\<spacespace> G,s\<turnstile>l [\<sim>\<Colon>\<preceq>]L\<spacespace> \<and>
--- a/src/HOL/Bali/Decl.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/Decl.thy Mon Mar 01 13:40:23 2010 +0100
@@ -206,8 +206,7 @@
"mdecl" <= (type) "sig \<times> methd"
-constdefs
- mhead::"methd \<Rightarrow> mhead"
+definition mhead :: "methd \<Rightarrow> mhead" where
"mhead m \<equiv> \<lparr>access=access m, static=static m, pars=pars m, resT=resT m\<rparr>"
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 \<times> memberdecl \<Rightarrow> bool"
+definition is_field :: "qtname \<times> memberdecl \<Rightarrow> bool" where
"is_field m \<equiv> \<exists> declC f. m=(declC,fdecl f)"
lemma is_fieldD: "is_field m \<Longrightarrow> \<exists> 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 \<times> memberdecl \<Rightarrow> bool"
+definition is_method :: "qtname \<times> memberdecl \<Rightarrow> bool" where
"is_method membr \<equiv> \<exists> declC m. membr=(declC,mdecl m)"
lemma is_methodD: "is_method membr \<Longrightarrow> \<exists> declC m. membr=(declC,mdecl m)"
@@ -315,8 +314,7 @@
isuperIfs::qtname list,\<dots>::'a\<rparr>"
"idecl" <= (type) "qtname \<times> iface"
-constdefs
- ibody :: "iface \<Rightarrow> ibody"
+definition ibody :: "iface \<Rightarrow> ibody" where
"ibody i \<equiv> \<lparr>access=access i,imethods=imethods i\<rparr>"
lemma access_ibody [simp]: "(access (ibody i)) = access i"
@@ -351,8 +349,7 @@
super::qtname,superIfs::qtname list,\<dots>::'a\<rparr>"
"cdecl" <= (type) "qtname \<times> class"
-constdefs
- cbody :: "class \<Rightarrow> cbody"
+definition cbody :: "class \<Rightarrow> cbody" where
"cbody c \<equiv> \<lparr>access=access c, cfields=cfields c,methods=methods c,init=init c\<rparr>"
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 \<equiv> [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) \<in>(subcls1 G)^+"
notation (xsymbols)
- subcls1_syntax ("_\<turnstile>_\<prec>\<^sub>C\<^sub>1_" [71,71,71] 70) and
+ subcls1_syntax ("_\<turnstile>_\<prec>\<^sub>C1_" [71,71,71] 70) and
subclseq_syntax ("_\<turnstile>_\<preceq>\<^sub>C _" [71,71,71] 70) and
subcls_syntax ("_\<turnstile>_\<prec>\<^sub>C _" [71,71,71] 70)
@@ -510,7 +507,7 @@
"\<lbrakk>G\<turnstile>C \<prec>\<^sub>C D\<rbrakk> \<Longrightarrow> \<exists> c. class G C = Some c"
by (auto simp add: subcls1_def dest: tranclD)
-lemma no_subcls1_Object:"G\<turnstile>Object\<prec>\<^sub>C\<^sub>1 D \<Longrightarrow> P"
+lemma no_subcls1_Object:"G\<turnstile>Object\<prec>\<^sub>C1 D \<Longrightarrow> P"
by (auto simp add: subcls1_def)
lemma no_subcls_Object: "G\<turnstile>Object\<prec>\<^sub>C D \<Longrightarrow> P"
@@ -520,14 +517,13 @@
section "well-structured programs"
-constdefs
- ws_idecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname list \<Rightarrow> bool"
+definition ws_idecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname list \<Rightarrow> bool" where
"ws_idecl G I si \<equiv> \<forall>J\<in>set si. is_iface G J \<and> (J,I)\<notin>(subint1 G)^+"
- ws_cdecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
+definition ws_cdecl :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" where
"ws_cdecl G C sc \<equiv> C\<noteq>Object \<longrightarrow> is_class G sc \<and> (sc,C)\<notin>(subcls1 G)^+"
- ws_prog :: "prog \<Rightarrow> bool"
+definition ws_prog :: "prog \<Rightarrow> bool" where
"ws_prog G \<equiv> (\<forall>(I,i)\<in>set (ifaces G). ws_idecl G I (isuperIfs i)) \<and>
(\<forall>(C,c)\<in>set (classes G). ws_cdecl G C (super c))"
@@ -680,7 +676,7 @@
then have "is_class G C \<Longrightarrow> P C"
proof (induct rule: subcls1_induct)
fix C
- assume hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<longrightarrow> is_class G S \<longrightarrow> P S"
+ assume hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C1 S \<longrightarrow> is_class G S \<longrightarrow> P S"
and iscls:"is_class G C"
show "P C"
proof (cases "C=Object")
@@ -715,7 +711,7 @@
then have "\<And> c. class G C = Some c\<Longrightarrow> P C c"
proof (induct rule: subcls1_induct)
fix C c
- assume hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<longrightarrow> (\<forall> s. class G S = Some s \<longrightarrow> P S s)"
+ assume hyp:"\<forall> S. G\<turnstile>C \<prec>\<^sub>C1 S \<longrightarrow> (\<forall> s. class G S = Some s \<longrightarrow> 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\<turnstile>C \<prec>\<^sub>C\<^sub>1 (super c)" by (rule subcls1I)
+ from iscls False have "G\<turnstile>C \<prec>\<^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 \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
+definition imethds :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables" where
--{* methods of an interface, with overriding and inheritance, cf. 9.2 *}
"imethds G I
\<equiv> iface_rec (G,I)
--- a/src/HOL/Bali/DeclConcepts.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/DeclConcepts.thy Mon Mar 01 13:40:23 2010 +0100
@@ -8,8 +8,7 @@
section "access control (cf. 6.6), overriding and hiding (cf. 8.4.6.1)"
-constdefs
-is_public :: "prog \<Rightarrow> qtname \<Rightarrow> bool"
+definition is_public :: "prog \<Rightarrow> qtname \<Rightarrow> bool" where
"is_public G qn \<equiv> (case class G qn of
None \<Rightarrow> (case iface G qn of
None \<Rightarrow> False
@@ -38,14 +37,16 @@
declare accessible_in_RefT_simp [simp del]
-constdefs
- is_acc_class :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
+definition is_acc_class :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool" where
"is_acc_class G P C \<equiv> is_class G C \<and> G\<turnstile>(Class C) accessible_in P"
- is_acc_iface :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool"
+
+definition is_acc_iface :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> bool" where
"is_acc_iface G P I \<equiv> is_iface G I \<and> G\<turnstile>(Iface I) accessible_in P"
- is_acc_type :: "prog \<Rightarrow> pname \<Rightarrow> ty \<Rightarrow> bool"
+
+definition is_acc_type :: "prog \<Rightarrow> pname \<Rightarrow> ty \<Rightarrow> bool" where
"is_acc_type G P T \<equiv> is_type G T \<and> G\<turnstile>T accessible_in P"
- is_acc_reftype :: "prog \<Rightarrow> pname \<Rightarrow> ref_ty \<Rightarrow> bool"
+
+definition is_acc_reftype :: "prog \<Rightarrow> pname \<Rightarrow> ref_ty \<Rightarrow> bool" where
"is_acc_reftype G P T \<equiv> isrtype G T \<and> G\<turnstile>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 \<times> mdecl) \<Rightarrow> (qtname \<times> memberdecl)"
+definition methdMembr :: "(qtname \<times> mdecl) \<Rightarrow> (qtname \<times> memberdecl)" where
"methdMembr m \<equiv> (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 \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> memberdecl)"
+definition method :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> memberdecl)" where
"method sig m \<equiv> (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 \<Rightarrow> (qtname \<times> field) \<Rightarrow> (qtname \<times> memberdecl)"
+definition fieldm :: "vname \<Rightarrow> (qtname \<times> field) \<Rightarrow> (qtname \<times> memberdecl)" where
"fieldm n f \<equiv> (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 \<times> mdecl) \<Rightarrow> sig"
+definition msig :: "(qtname \<times> mdecl) \<Rightarrow> sig" where
"msig m \<equiv> 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 \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> mdecl)"
+definition qmdecl :: "sig \<Rightarrow> (qtname \<times> methd) \<Rightarrow> (qtname \<times> mdecl)" where
"qmdecl sig m \<equiv> (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 \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool"
- ("_ \<turnstile> _ inheritable'_in _" [61,61,61] 60)
+definition inheritable_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> pname \<Rightarrow> bool" ("_ \<turnstile> _ inheritable'_in _" [61,61,61] 60) where
+
"G\<turnstile>membr inheritable_in pack
\<equiv> (case (accmodi membr) of
Private \<Rightarrow> False
@@ -529,25 +525,21 @@
subsubsection "declared-in/undeclared-in"
-constdefs cdeclaredmethd:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,methd) table"
+definition cdeclaredmethd :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,methd) table" where
"cdeclaredmethd G C
\<equiv> (case class G C of
None \<Rightarrow> \<lambda> sig. None
| Some c \<Rightarrow> table_of (methods c)
)"
-constdefs
-cdeclaredfield:: "prog \<Rightarrow> qtname \<Rightarrow> (vname,field) table"
+definition cdeclaredfield :: "prog \<Rightarrow> qtname \<Rightarrow> (vname,field) table" where
"cdeclaredfield G C
\<equiv> (case class G C of
None \<Rightarrow> \<lambda> sig. None
| Some c \<Rightarrow> table_of (cfields c)
)"
-
-constdefs
-declared_in:: "prog \<Rightarrow> memberdecl \<Rightarrow> qtname \<Rightarrow> bool"
- ("_\<turnstile> _ declared'_in _" [61,61,61] 60)
+definition declared_in :: "prog \<Rightarrow> memberdecl \<Rightarrow> qtname \<Rightarrow> bool" ("_\<turnstile> _ declared'_in _" [61,61,61] 60) where
"G\<turnstile>m declared_in C \<equiv> (case m of
fdecl (fn,f ) \<Rightarrow> cdeclaredfield G C fn = Some f
| mdecl (sig,m) \<Rightarrow> 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 \<Rightarrow> memberid \<Rightarrow> qtname \<Rightarrow> bool"
- ("_\<turnstile> _ undeclared'_in _" [61,61,61] 60)
-
+definition undeclared_in :: "prog \<Rightarrow> memberid \<Rightarrow> qtname \<Rightarrow> bool" ("_\<turnstile> _ undeclared'_in _" [61,61,61] 60) where
"G\<turnstile>m undeclared_in C \<equiv> (case m of
fid fn \<Rightarrow> cdeclaredfield G C fn = None
| mid sig \<Rightarrow> cdeclaredmethd G C sig = None)"
@@ -591,7 +580,7 @@
Immediate: "\<lbrakk>G\<turnstile>mbr m declared_in C;declclass m = C\<rbrakk> \<Longrightarrow> G\<turnstile>m member_of C"
| Inherited: "\<lbrakk>G\<turnstile>m inheritable_in (pid C); G\<turnstile>memberid m undeclared_in C;
- G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S
+ G\<turnstile>C \<prec>\<^sub>C1 S; G\<turnstile>(Class S) accessible_in (pid C);G\<turnstile>m member_of S
\<rbrakk> \<Longrightarrow> G\<turnstile>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 @@
("_ \<turnstile>Field _ _ member'_of _" [61,61,61] 60)
where "G\<turnstile>Field n f member_of C == G\<turnstile>fieldm n f member_of C"
-constdefs
-inherits:: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> bool"
- ("_ \<turnstile> _ inherits _" [61,61,61] 60)
+definition inherits :: "prog \<Rightarrow> qtname \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> bool" ("_ \<turnstile> _ inherits _" [61,61,61] 60) where
"G\<turnstile>C inherits m
\<equiv> G\<turnstile>m inheritable_in (pid C) \<and> G\<turnstile>memberid m undeclared_in C \<and>
- (\<exists> S. G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S \<and> G\<turnstile>(Class S) accessible_in (pid C) \<and> G\<turnstile>m member_of S)"
+ (\<exists> S. G\<turnstile>C \<prec>\<^sub>C1 S \<and> G\<turnstile>(Class S) accessible_in (pid C) \<and> G\<turnstile>m member_of S)"
lemma inherits_member: "G\<turnstile>C inherits m \<Longrightarrow> G\<turnstile>m member_of C"
by (auto simp add: inherits_def intro: members.Inherited)
-constdefs member_in::"prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile> _ member'_in _" [61,61,61] 60)
+definition member_in :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> bool" ("_ \<turnstile> _ member'_in _" [61,61,61] 60) where
"G\<turnstile>m member_in C \<equiv> \<exists> provC. G\<turnstile> C \<preceq>\<^sub>C provC \<and> G \<turnstile> 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\<turnstile>Method new declared_in (declclass new);
G\<turnstile>Method old declared_in (declclass old);
G\<turnstile>Method old inheritable_in pid (declclass new);
- G\<turnstile>(declclass new) \<prec>\<^sub>C\<^sub>1 superNew;
+ G\<turnstile>(declclass new) \<prec>\<^sub>C1 superNew;
G \<turnstile>Method old member_of superNew
\<rbrakk> \<Longrightarrow> G\<turnstile>new overrides\<^sub>S old"
@@ -716,9 +702,7 @@
subsubsection "Hiding"
-constdefs hides::
-"prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool"
- ("_\<turnstile> _ hides _" [61,61,61] 60)
+definition hides :: "prog \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> (qtname \<times> mdecl) \<Rightarrow> bool" ("_\<turnstile> _ hides _" [61,61,61] 60) where
"G\<turnstile>new hides old
\<equiv> is_static new \<and> msig new = msig old \<and>
G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old) \<and>
@@ -777,11 +761,7 @@
by (auto simp add: hides_def)
subsubsection "permits access"
-constdefs
-permits_acc::
- "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool"
- ("_ \<turnstile> _ in _ permits'_acc'_from _" [61,61,61,61] 60)
-
+definition permits_acc :: "prog \<Rightarrow> (qtname \<times> memberdecl) \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> bool" ("_ \<turnstile> _ in _ permits'_acc'_from _" [61,61,61,61] 60) where
"G\<turnstile>membr in cls permits_acc_from accclass
\<equiv> (case (accmodi membr) of
Private \<Rightarrow> (declclass membr = accclass)
@@ -980,7 +960,7 @@
next
case (Inherited n C S)
assume undecl: "G\<turnstile> memberid n undeclared_in C"
- assume super: "G\<turnstile>C\<prec>\<^sub>C\<^sub>1S"
+ assume super: "G\<turnstile>C\<prec>\<^sub>C1S"
assume hyp: "\<lbrakk>G \<turnstile> m member_of S; memberid n = memberid m\<rbrakk> \<Longrightarrow> n = m"
assume eqid: "memberid n = memberid m"
assume "G \<turnstile> 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\<turnstile>C\<prec>\<^sub>C\<^sub>1S" and "is_class G S"
+ assume "G\<turnstile>C\<prec>\<^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:
-"\<lbrakk>G\<turnstile>(declclass new) \<prec>\<^sub>C\<^sub>1 superNew; G \<turnstile>Method old member_of superNew \<rbrakk>
+"\<lbrakk>G\<turnstile>(declclass new) \<prec>\<^sub>C1 superNew; G \<turnstile>Method old member_of superNew \<rbrakk>
\<Longrightarrow> G\<turnstile>(declclass new) \<prec>\<^sub>C (declclass old)"
apply (rule trancl_rtrancl_trancl)
apply (erule r_into_trancl)
@@ -1257,7 +1237,7 @@
"G\<turnstile> memberid m undeclared_in D"
"G \<turnstile> Class S accessible_in pid D"
"G \<turnstile> m member_of S"
- assume super: "G\<turnstile>D\<prec>\<^sub>C\<^sub>1S"
+ assume super: "G\<turnstile>D\<prec>\<^sub>C1S"
assume hyp: "\<lbrakk>G\<turnstile>S\<preceq>\<^sub>C C; G\<turnstile>C\<preceq>\<^sub>C declclass m\<rbrakk> \<Longrightarrow> G \<turnstile> m member_of C"
assume subclseq_C_m: "G\<turnstile>C\<preceq>\<^sub>C declclass m"
assume "G\<turnstile>D\<preceq>\<^sub>C C"
@@ -1399,24 +1379,21 @@
translations
"fspec" <= (type) "vname \<times> qtname"
-constdefs
-imethds:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
+definition imethds :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables" where
"imethds G I
\<equiv> iface_rec (G,I)
(\<lambda>I i ts. (Un_tables ts) \<oplus>\<oplus>
(Option.set \<circ> table_of (map (\<lambda>(s,m). (s,I,m)) (imethods i))))"
text {* methods of an interface, with overriding and inheritance, cf. 9.2 *}
-constdefs
-accimethds:: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables"
+definition accimethds :: "prog \<Rightarrow> pname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> mhead) tables" where
"accimethds G pack I
\<equiv> if G\<turnstile>Iface I accessible_in pack
then imethds G I
else \<lambda> k. {}"
text {* only returns imethds if the interface is accessible *}
-constdefs
-methd:: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+definition methd :: "prog \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
"methd G C
\<equiv> 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 \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+definition accmethd :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
"accmethd G S C
\<equiv> filter_tab (\<lambda>sig m. G\<turnstile>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 \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+definition dynmethd :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
"dynmethd G statC dynC
\<equiv> \<lambda> sig.
(if G\<turnstile>dynC \<preceq>\<^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 \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+definition dynimethd :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
"dynimethd G I dynC
\<equiv> \<lambda> sig. if imethds G I sig \<noteq> {}
then methd G dynC sig
@@ -1493,8 +1467,7 @@
down to the actual dynamic class.
*}
-constdefs
-dynlookup::"prog \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table"
+definition dynlookup :: "prog \<Rightarrow> ref_ty \<Rightarrow> qtname \<Rightarrow> (sig,qtname \<times> methd) table" where
"dynlookup G statT dynC
\<equiv> (case statT of
NullT \<Rightarrow> 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 \<Rightarrow> qtname \<Rightarrow> ((vname \<times> qtname) \<times> field) list"
+definition fields :: "prog \<Rightarrow> qtname \<Rightarrow> ((vname \<times> qtname) \<times> field) list" where
"fields G C
\<equiv> class_rec (G,C) [] (\<lambda>C c ts. map (\<lambda>(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 \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (vname, qtname \<times> field) table"
+definition accfield :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> (vname, qtname \<times> field) table" where
"accfield G S C
\<equiv> let field_tab = table_of((map (\<lambda>((n,d),f).(n,(d,f)))) (fields G C))
in filter_tab (\<lambda>n (declC,f). G\<turnstile> (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 \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> bool"
+definition is_methd :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> bool" where
"is_methd G \<equiv> \<lambda>C sig. is_class G C \<and> methd G C sig \<noteq> None"
-constdefs efname:: "((vname \<times> qtname) \<times> field) \<Rightarrow> (vname \<times> qtname)"
+definition efname :: "((vname \<times> qtname) \<times> field) \<Rightarrow> (vname \<times> qtname)" where
"efname \<equiv> fst"
lemma efname_simp[simp]:"efname (n,f) = n"
@@ -2300,8 +2269,7 @@
section "calculation of the superclasses of a class"
-constdefs
- superclasses:: "prog \<Rightarrow> qtname \<Rightarrow> qtname set"
+definition superclasses :: "prog \<Rightarrow> qtname \<Rightarrow> qtname set" where
"superclasses G C \<equiv> class_rec (G,C) {}
(\<lambda> C c superclss. (if C=Object
then {}
--- a/src/HOL/Bali/DefiniteAssignment.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/DefiniteAssignment.thy Mon Mar 01 13:40:23 2010 +0100
@@ -74,7 +74,7 @@
"jumpNestingOkS jmps (FinA a c) = False"
-constdefs jumpNestingOk :: "jump set \<Rightarrow> term \<Rightarrow> bool"
+definition jumpNestingOk :: "jump set \<Rightarrow> term \<Rightarrow> bool" where
"jumpNestingOk jmps t \<equiv> (case t of
In1 se \<Rightarrow> (case se of
Inl e \<Rightarrow> True
@@ -156,7 +156,7 @@
"assignsEs [] = {}"
"assignsEs (e#es) = assignsE e \<union> assignsEs es"
-constdefs assigns:: "term \<Rightarrow> lname set"
+definition assigns :: "term \<Rightarrow> lname set" where
"assigns t \<equiv> (case t of
In1 se \<Rightarrow> (case se of
Inl e \<Rightarrow> assignsE e
@@ -429,20 +429,14 @@
subsection {* Lifting set operations to range of tables (map to a set) *}
-constdefs
- union_ts:: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
- ("_ \<Rightarrow>\<union> _" [67,67] 65)
+definition union_ts :: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" ("_ \<Rightarrow>\<union> _" [67,67] 65) where
"A \<Rightarrow>\<union> B \<equiv> \<lambda> k. A k \<union> B k"
-constdefs
- intersect_ts:: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
- ("_ \<Rightarrow>\<inter> _" [72,72] 71)
+definition intersect_ts :: "('a,'b) tables \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" ("_ \<Rightarrow>\<inter> _" [72,72] 71) where
"A \<Rightarrow>\<inter> B \<equiv> \<lambda> k. A k \<inter> B k"
-constdefs
- all_union_ts:: "('a,'b) tables \<Rightarrow> 'b set \<Rightarrow> ('a,'b) tables"
- (infixl "\<Rightarrow>\<union>\<^sub>\<forall>" 40)
-"A \<Rightarrow>\<union>\<^sub>\<forall> B \<equiv> \<lambda> k. A k \<union> B"
+definition all_union_ts :: "('a,'b) tables \<Rightarrow> 'b set \<Rightarrow> ('a,'b) tables" (infixl "\<Rightarrow>\<union>\<^sub>\<forall>" 40) where
+ "A \<Rightarrow>\<union>\<^sub>\<forall> B \<equiv> \<lambda> k. A k \<union> B"
subsubsection {* Binary union of tables *}
@@ -513,15 +507,15 @@
brk :: "breakass" --{* Definetly assigned variables for
abrupt completion with a break *}
-constdefs rmlab :: "'a \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables"
+definition rmlab :: "'a \<Rightarrow> ('a,'b) tables \<Rightarrow> ('a,'b) tables" where
"rmlab k A \<equiv> \<lambda> x. if x=k then UNIV else A x"
(*
-constdefs setbrk :: "breakass \<Rightarrow> assigned \<Rightarrow> breakass set"
+definition setbrk :: "breakass \<Rightarrow> assigned \<Rightarrow> breakass set" where
"setbrk b A \<equiv> {b} \<union> {a| a. a\<in> brk A \<and> lab a \<noteq> lab b}"
*)
-constdefs range_inter_ts :: "('a,'b) tables \<Rightarrow> 'b set" ("\<Rightarrow>\<Inter>_" 80)
+definition range_inter_ts :: "('a,'b) tables \<Rightarrow> 'b set" ("\<Rightarrow>\<Inter>_" 80) where
"\<Rightarrow>\<Inter>A \<equiv> {x |x. \<forall> k. x \<in> A k}"
text {*
--- a/src/HOL/Bali/Eval.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/Eval.thy Mon Mar 01 13:40:23 2010 +0100
@@ -140,8 +140,7 @@
lst_inj_vals ("\<lfloor>_\<rfloor>\<^sub>l" 1000)
where "\<lfloor>es\<rfloor>\<^sub>l == In3 es"
-constdefs
- undefined3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals"
+definition undefined3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals" where
"undefined3 \<equiv> sum3_case (In1 \<circ> sum_case (\<lambda>x. undefined) (\<lambda>x. Unit))
(\<lambda>x. In2 undefined) (\<lambda>x. In3 undefined)"
@@ -160,8 +159,7 @@
section "exception throwing and catching"
-constdefs
- throw :: "val \<Rightarrow> abopt \<Rightarrow> abopt"
+definition throw :: "val \<Rightarrow> abopt \<Rightarrow> abopt" where
"throw a' x \<equiv> 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 \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60)
+definition fits :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60) where
"G,s\<turnstile>a' fits T \<equiv> (\<exists>rt. T=RefT rt) \<longrightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"
lemma fits_Null [simp]: "G,s\<turnstile>Null fits T"
@@ -195,8 +192,7 @@
apply iprover
done
-constdefs
- catch ::"prog \<Rightarrow> state \<Rightarrow> qtname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60)
+definition catch :: "prog \<Rightarrow> state \<Rightarrow> qtname \<Rightarrow> bool" ("_,_\<turnstile>catch _"[61,61,61]60) where
"G,s\<turnstile>catch C\<equiv>\<exists>xc. abrupt s=Some (Xcpt xc) \<and>
G,store s\<turnstile>Addr (the_Loc xc) fits Class C"
@@ -221,8 +217,7 @@
apply (simp (no_asm))
done
-constdefs
- new_xcpt_var :: "vname \<Rightarrow> state \<Rightarrow> state"
+definition new_xcpt_var :: "vname \<Rightarrow> state \<Rightarrow> state" where
"new_xcpt_var vn \<equiv>
\<lambda>(x,s). Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)"
@@ -237,9 +232,7 @@
section "misc"
-constdefs
-
- assign :: "('a \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> 'a \<Rightarrow> state \<Rightarrow> state"
+definition assign :: "('a \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> 'a \<Rightarrow> state \<Rightarrow> state" where
"assign f v \<equiv> \<lambda>(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 \<Rightarrow> stmt"
+definition init_comp_ty :: "ty \<Rightarrow> stmt" where
"init_comp_ty T \<equiv> if (\<exists>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 \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname"
+definition invocation_class :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname" where
"invocation_class m s a' statT
\<equiv> (case m of
Static \<Rightarrow> if (\<exists> statC. statT = ClassT statC)
@@ -321,7 +310,7 @@
| SuperM \<Rightarrow> the_Class (RefT statT)
| IntVir \<Rightarrow> obj_class (lookup_obj s a'))"
-invocation_declclass::"prog \<Rightarrow> inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> qtname"
+definition invocation_declclass::"prog \<Rightarrow> inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> qtname" where
"invocation_declclass G m s a' statT sig
\<equiv> 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 \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> inv_mode \<Rightarrow> val \<Rightarrow> val list \<Rightarrow>
- state \<Rightarrow> state"
+definition init_lvars :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> inv_mode \<Rightarrow> val \<Rightarrow> val list \<Rightarrow>
+ state \<Rightarrow> state" where
"init_lvars G C sig mode a' pvs
\<equiv> \<lambda> (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 \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> expr"
+definition body :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> expr" where
"body G C sig \<equiv> let m = the (methd G C sig)
in Body (declclass m) (stmt (mbody (mthd m)))"
@@ -390,12 +377,10 @@
section "variables"
-constdefs
-
- lvar :: "lname \<Rightarrow> st \<Rightarrow> vvar"
+definition lvar :: "lname \<Rightarrow> st \<Rightarrow> vvar" where
"lvar vn s \<equiv> (the (locals s vn), \<lambda>v. supd (lupd(vn\<mapsto>v)))"
- fvar :: "qtname \<Rightarrow> bool \<Rightarrow> vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
+definition fvar :: "qtname \<Rightarrow> bool \<Rightarrow> vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state" where
"fvar C stat fn a' s
\<equiv> let (oref,xf) = if stat then (Stat C,id)
else (Heap (the_Addr a'),np a');
@@ -403,7 +388,7 @@
f = (\<lambda>v. supd (upd_gobj oref n v))
in ((the (values (the (globs (store s) oref)) n),f),abupd xf s)"
- avar :: "prog \<Rightarrow> val \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
+definition avar :: "prog \<Rightarrow> val \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state" where
"avar G i' a' s
\<equiv> 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 \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> vname \<Rightarrow> bool \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state"
+definition check_field_access :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> vname \<Rightarrow> bool \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state" where
"check_field_access G accC statDeclC fn stat a' s
\<equiv> let oref = if stat then Stat statDeclC
else Heap (the_Addr a');
@@ -461,9 +444,7 @@
AccessViolation)
s"
-constdefs
-check_method_access::
- "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> inv_mode \<Rightarrow> sig \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state"
+definition check_method_access :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> inv_mode \<Rightarrow> sig \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state" where
"check_method_access G accC statT mode sig a' s
\<equiv> let invC = invocation_class mode (store s) a' statT;
dynM = the (dynlookup G statT invC sig)
--- a/src/HOL/Bali/Example.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/Example.thy Mon Mar 01 13:40:23 2010 +0100
@@ -153,23 +153,18 @@
foo :: mname
-constdefs
-
- foo_sig :: sig
- "foo_sig \<equiv> \<lparr>name=foo,parTs=[Class Base]\<rparr>"
+definition foo_sig :: sig
+ where "foo_sig \<equiv> \<lparr>name=foo,parTs=[Class Base]\<rparr>"
- foo_mhead :: mhead
- "foo_mhead \<equiv> \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>"
+definition foo_mhead :: mhead
+ where "foo_mhead \<equiv> \<lparr>access=Public,static=False,pars=[z],resT=Class Base\<rparr>"
-constdefs
-
- Base_foo :: mdecl
- "Base_foo \<equiv> (foo_sig, \<lparr>access=Public,static=False,pars=[z],resT=Class Base,
+definition Base_foo :: mdecl
+ where "Base_foo \<equiv> (foo_sig, \<lparr>access=Public,static=False,pars=[z],resT=Class Base,
mbody=\<lparr>lcls=[],stmt=Return (!!z)\<rparr>\<rparr>)"
-constdefs
- Ext_foo :: mdecl
- "Ext_foo \<equiv> (foo_sig,
+definition Ext_foo :: mdecl
+ where "Ext_foo \<equiv> (foo_sig,
\<lparr>access=Public,static=False,pars=[z],resT=Class Ext,
mbody=\<lparr>lcls=[]
,stmt=Expr({Ext,Ext,False}Cast (Class Ext) (!!z)..vee :=
@@ -177,12 +172,10 @@
Return (Lit Null)\<rparr>
\<rparr>)"
-constdefs
-
-arr_viewed_from :: "qtname \<Rightarrow> qtname \<Rightarrow> var"
+definition arr_viewed_from :: "qtname \<Rightarrow> qtname \<Rightarrow> var" where
"arr_viewed_from accC C \<equiv> {accC,Base,True}StatRef (ClassT C)..arr"
-BaseCl :: "class"
+definition BaseCl :: "class" where
"BaseCl \<equiv> \<lparr>access=Public,
cfields=[(arr, \<lparr>access=Public,static=True ,type=PrimT Boolean.[]\<rparr>),
(vee, \<lparr>access=Public,static=False,type=Iface HasFoo \<rparr>)],
@@ -192,7 +185,7 @@
super=Object,
superIfs=[HasFoo]\<rparr>"
-ExtCl :: "class"
+definition ExtCl :: "class" where
"ExtCl \<equiv> \<lparr>access=Public,
cfields=[(vee, \<lparr>access=Public,static=False,type= PrimT Integer\<rparr>)],
methods=[Ext_foo],
@@ -200,7 +193,7 @@
super=Base,
superIfs=[]\<rparr>"
-MainCl :: "class"
+definition MainCl :: "class" where
"MainCl \<equiv> \<lparr>access=Public,
cfields=[],
methods=[],
@@ -209,16 +202,14 @@
superIfs=[]\<rparr>"
(* The "main" method is modeled seperately (see tprg) *)
-constdefs
-
- HasFooInt :: iface
- "HasFooInt \<equiv> \<lparr>access=Public,imethods=[(foo_sig, foo_mhead)],isuperIfs=[]\<rparr>"
+definition HasFooInt :: iface
+ where "HasFooInt \<equiv> \<lparr>access=Public,imethods=[(foo_sig, foo_mhead)],isuperIfs=[]\<rparr>"
- Ifaces ::"idecl list"
- "Ifaces \<equiv> [(HasFoo,HasFooInt)]"
+definition Ifaces ::"idecl list"
+ where "Ifaces \<equiv> [(HasFoo,HasFooInt)]"
- "Classes" ::"cdecl list"
- "Classes \<equiv> [(Base,BaseCl),(Ext,ExtCl),(Main,MainCl)]@standard_classes"
+definition "Classes" ::"cdecl list"
+ where "Classes \<equiv> [(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 == \<lparr>ifaces=Ifaces,classes=Classes\<rparr>"
-constdefs
- test :: "(ty)list \<Rightarrow> stmt"
+definition test :: "(ty)list \<Rightarrow> stmt" where
"test pTs \<equiv> e:==NewC Ext;;
\<spacespace> Try Expr({Main,ClassT Base,IntVir}!!e\<cdot>foo({pTs}[Lit Null]))
\<spacespace> Catch((SXcpt NullPointer) z)
--- a/src/HOL/Bali/State.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/State.thy Mon Mar 01 13:40:23 2010 +0100
@@ -38,9 +38,7 @@
"obj" <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option\<rparr>"
"obj" <= (type) "\<lparr>tag::obj_tag, values::vn \<Rightarrow> val option,\<dots>::'a\<rparr>"
-constdefs
-
- the_Arr :: "obj option \<Rightarrow> ty \<times> int \<times> (vn, val) table"
+definition the_Arr :: "obj option \<Rightarrow> ty \<times> int \<times> (vn, val) table" where
"the_Arr obj \<equiv> SOME (T,k,t). obj = Some \<lparr>tag=Arr T k,values=t\<rparr>"
lemma the_Arr_Arr [simp]: "the_Arr (Some \<lparr>tag=Arr T k,values=cs\<rparr>) = (T,k,cs)"
@@ -52,9 +50,7 @@
apply (auto simp add: the_Arr_def)
done
-constdefs
-
- upd_obj :: "vn \<Rightarrow> val \<Rightarrow> obj \<Rightarrow> obj"
+definition upd_obj :: "vn \<Rightarrow> val \<Rightarrow> obj \<Rightarrow> obj" where
"upd_obj n v \<equiv> \<lambda> obj . obj \<lparr>values:=(values obj)(n\<mapsto>v)\<rparr>"
lemma upd_obj_def2 [simp]:
@@ -62,8 +58,7 @@
apply (auto simp: upd_obj_def)
done
-constdefs
- obj_ty :: "obj \<Rightarrow> ty"
+definition obj_ty :: "obj \<Rightarrow> ty" where
"obj_ty obj \<equiv> case tag obj of
CInst C \<Rightarrow> Class C
| Arr T k \<Rightarrow> T.[]"
@@ -102,9 +97,7 @@
apply (auto split add: obj_tag.split_asm)
done
-constdefs
-
- obj_class :: "obj \<Rightarrow> qtname"
+definition obj_class :: "obj \<Rightarrow> qtname" where
"obj_class obj \<equiv> case tag obj of
CInst C \<Rightarrow> C
| Arr T k \<Rightarrow> Object"
@@ -143,9 +136,7 @@
"Stat" => "CONST Inr"
"oref" <= (type) "loc + qtname"
-constdefs
- fields_table::
- "prog \<Rightarrow> qtname \<Rightarrow> (fspec \<Rightarrow> field \<Rightarrow> bool) \<Rightarrow> (fspec, ty) table"
+definition fields_table :: "prog \<Rightarrow> qtname \<Rightarrow> (fspec \<Rightarrow> field \<Rightarrow> bool) \<Rightarrow> (fspec, ty) table" where
"fields_table G C P
\<equiv> Option.map type \<circ> table_of (filter (split P) (DeclConcepts.fields G C))"
@@ -182,14 +173,13 @@
apply simp
done
-constdefs
- in_bounds :: "int \<Rightarrow> int \<Rightarrow> bool" ("(_/ in'_bounds _)" [50, 51] 50)
+definition in_bounds :: "int \<Rightarrow> int \<Rightarrow> bool" ("(_/ in'_bounds _)" [50, 51] 50) where
"i in_bounds k \<equiv> 0 \<le> i \<and> i < k"
- arr_comps :: "'a \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a option"
+definition arr_comps :: "'a \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a option" where
"arr_comps T k \<equiv> \<lambda>i. if i in_bounds k then Some T else None"
- var_tys :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> (vn, ty) table"
+definition var_tys :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> (vn, ty) table" where
"var_tys G oi r
\<equiv> case r of
Heap a \<Rightarrow> (case oi of
@@ -232,15 +222,13 @@
subsection "access"
-constdefs
-
- globs :: "st \<Rightarrow> globs"
+definition globs :: "st \<Rightarrow> globs" where
"globs \<equiv> st_case (\<lambda>g l. g)"
- locals :: "st \<Rightarrow> locals"
+definition locals :: "st \<Rightarrow> locals" where
"locals \<equiv> st_case (\<lambda>g l. l)"
- heap :: "st \<Rightarrow> heap"
+definition heap :: "st \<Rightarrow> heap" where
"heap s \<equiv> globs s \<circ> Heap"
@@ -262,8 +250,7 @@
subsection "memory allocation"
-constdefs
- new_Addr :: "heap \<Rightarrow> loc option"
+definition new_Addr :: "heap \<Rightarrow> loc option" where
"new_Addr h \<equiv> if (\<forall>a. h a \<noteq> None) then None else Some (SOME a. h a = None)"
lemma new_AddrD: "new_Addr h = Some a \<Longrightarrow> h a = None"
@@ -303,20 +290,19 @@
subsection "update"
-constdefs
- gupd :: "oref \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st" ("gupd'(_\<mapsto>_')"[10,10]1000)
+definition gupd :: "oref \<Rightarrow> obj \<Rightarrow> st \<Rightarrow> st" ("gupd'(_\<mapsto>_')"[10,10]1000) where
"gupd r obj \<equiv> st_case (\<lambda>g l. st (g(r\<mapsto>obj)) l)"
- lupd :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" ("lupd'(_\<mapsto>_')"[10,10]1000)
+definition lupd :: "lname \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" ("lupd'(_\<mapsto>_')"[10,10]1000) where
"lupd vn v \<equiv> st_case (\<lambda>g l. st g (l(vn\<mapsto>v)))"
- upd_gobj :: "oref \<Rightarrow> vn \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st"
+definition upd_gobj :: "oref \<Rightarrow> vn \<Rightarrow> val \<Rightarrow> st \<Rightarrow> st" where
"upd_gobj r n v \<equiv> st_case (\<lambda>g l. st (chg_map (upd_obj n v) r g) l)"
- set_locals :: "locals \<Rightarrow> st \<Rightarrow> st"
+definition set_locals :: "locals \<Rightarrow> st \<Rightarrow> st" where
"set_locals l \<equiv> st_case (\<lambda>g l'. st g l)"
- init_obj :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> st \<Rightarrow> st"
+definition init_obj :: "prog \<Rightarrow> obj_tag \<Rightarrow> oref \<Rightarrow> st \<Rightarrow> st" where
"init_obj G oi r \<equiv> gupd(r\<mapsto>\<lparr>tag=oi, values=init_vals (var_tys G oi r)\<rparr>)"
abbreviation
@@ -476,8 +462,7 @@
-constdefs
- abrupt_if :: "bool \<Rightarrow> abopt \<Rightarrow> abopt \<Rightarrow> abopt"
+definition abrupt_if :: "bool \<Rightarrow> abopt \<Rightarrow> abopt \<Rightarrow> abopt" where
"abrupt_if c x' x \<equiv> if c \<and> (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 \<Rightarrow> abopt \<Rightarrow> abopt"
+definition absorb :: "jump \<Rightarrow> abopt \<Rightarrow> abopt" where
"absorb j a \<equiv> if a=Some (Jump j) then None else a"
lemma absorb_SomeD [dest!]: "absorb j a = Some x \<Longrightarrow> a = Some x"
@@ -611,9 +595,7 @@
apply clarsimp
done
-constdefs
-
- normal :: "state \<Rightarrow> bool"
+definition normal :: "state \<Rightarrow> bool" where
"normal \<equiv> \<lambda>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 \<Rightarrow> state \<Rightarrow> bool"
+definition heap_free :: "nat \<Rightarrow> state \<Rightarrow> bool" where
"heap_free n \<equiv> \<lambda>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 \<Rightarrow> abopt) \<Rightarrow> state \<Rightarrow> state"
+definition abupd :: "(abopt \<Rightarrow> abopt) \<Rightarrow> state \<Rightarrow> state" where
"abupd f \<equiv> prod_fun f id"
- supd :: "(st \<Rightarrow> st) \<Rightarrow> state \<Rightarrow> state"
+definition supd :: "(st \<Rightarrow> st) \<Rightarrow> state \<Rightarrow> state" where
"supd \<equiv> prod_fun id"
lemma abupd_def2 [simp]: "abupd f (x,s) = (f x,s)"
@@ -692,12 +671,10 @@
section "initialisation test"
-constdefs
-
- inited :: "qtname \<Rightarrow> globs \<Rightarrow> bool"
+definition inited :: "qtname \<Rightarrow> globs \<Rightarrow> bool" where
"inited C g \<equiv> g (Stat C) \<noteq> None"
- initd :: "qtname \<Rightarrow> state \<Rightarrow> bool"
+definition initd :: "qtname \<Rightarrow> state \<Rightarrow> bool" where
"initd C \<equiv> inited C \<circ> globs \<circ> store"
lemma not_inited_empty [simp]: "\<not>inited C empty"
@@ -731,7 +708,7 @@
done
section {* @{text error_free} *}
-constdefs error_free:: "state \<Rightarrow> bool"
+definition error_free :: "state \<Rightarrow> bool" where
"error_free s \<equiv> \<not> (\<exists> err. abrupt s = Some (Error err))"
lemma error_free_Norm [simp,intro]: "error_free (Norm s)"
--- a/src/HOL/Bali/Table.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/Table.thy Mon Mar 01 13:40:23 2010 +0100
@@ -53,9 +53,7 @@
by (simp add: map_add_def)
section {* Conditional Override *}
-constdefs
-cond_override::
- "('b \<Rightarrow>'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b) table"
+definition cond_override :: "('b \<Rightarrow>'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b)table \<Rightarrow> ('a, 'b) table" where
--{* when merging tables old and new, only override an entry of table old when
the condition cond holds *}
@@ -100,8 +98,7 @@
section {* Filter on Tables *}
-constdefs
-filter_tab:: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) table \<Rightarrow> ('a, 'b) table"
+definition filter_tab :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a, 'b) table \<Rightarrow> ('a, 'b) table" where
"filter_tab c t \<equiv> \<lambda> k. (case t k of
None \<Rightarrow> None
| Some x \<Rightarrow> if c k x then Some x else None)"
--- a/src/HOL/Bali/Term.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/Term.thy Mon Mar 01 13:40:23 2010 +0100
@@ -261,9 +261,7 @@
StatRef :: "ref_ty \<Rightarrow> expr"
where "StatRef rt == Cast (RefT rt) (Lit Null)"
-constdefs
-
- is_stmt :: "term \<Rightarrow> bool"
+definition is_stmt :: "term \<Rightarrow> bool" where
"is_stmt t \<equiv> \<exists>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) \<and> (the_Bool v2))"
"eval_binop CondOr v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
-constdefs need_second_arg :: "binop \<Rightarrow> val \<Rightarrow> bool"
+definition need_second_arg :: "binop \<Rightarrow> val \<Rightarrow> bool" where
"need_second_arg binop v1 \<equiv> \<not> ((binop=CondAnd \<and> \<not> the_Bool v1) \<or>
(binop=CondOr \<and> the_Bool v1))"
text {* @{term CondAnd} and @{term CondOr} only evalulate the second argument
--- a/src/HOL/Bali/Trans.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/Trans.thy Mon Mar 01 13:40:23 2010 +0100
@@ -9,7 +9,7 @@
theory Trans imports Evaln begin
-constdefs groundVar:: "var \<Rightarrow> bool"
+definition groundVar :: "var \<Rightarrow> bool" where
"groundVar v \<equiv> (case v of
LVar ln \<Rightarrow> True
| {accC,statDeclC,stat}e..fn \<Rightarrow> \<exists> a. e=Lit a
@@ -34,7 +34,7 @@
done
qed
-constdefs groundExprs:: "expr list \<Rightarrow> bool"
+definition groundExprs :: "expr list \<Rightarrow> bool" where
"groundExprs es \<equiv> list_all (\<lambda> e. \<exists> v. e=Lit v) es"
consts the_val:: "expr \<Rightarrow> val"
--- a/src/HOL/Bali/Type.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/Type.thy Mon Mar 01 13:40:23 2010 +0100
@@ -41,8 +41,7 @@
abbreviation Array :: "ty \<Rightarrow> ty" ("_.[]" [90] 90)
where "T.[] == RefT (ArrayT T)"
-constdefs
- the_Class :: "ty \<Rightarrow> qtname"
+definition the_Class :: "ty \<Rightarrow> qtname" where
"the_Class T \<equiv> SOME C. T = Class C" (** primrec not possible here **)
lemma the_Class_eq [simp]: "the_Class (Class C)= C"
--- a/src/HOL/Bali/TypeRel.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/TypeRel.thy Mon Mar 01 13:40:23 2010 +0100
@@ -65,7 +65,7 @@
done
lemma subcls1I1:
- "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>\<^sub>C\<^sub>1 D"
+ "\<lbrakk>class G C = Some c; C \<noteq> Object;D=super c\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>\<^sub>C1 D"
apply (auto dest: subcls1I)
done
@@ -126,7 +126,7 @@
done
lemma single_inheritance:
-"\<lbrakk>G\<turnstile>A \<prec>\<^sub>C\<^sub>1 B; G\<turnstile>A \<prec>\<^sub>C\<^sub>1 C\<rbrakk> \<Longrightarrow> B = C"
+"\<lbrakk>G\<turnstile>A \<prec>\<^sub>C1 B; G\<turnstile>A \<prec>\<^sub>C1 C\<rbrakk> \<Longrightarrow> B = C"
by (auto simp add: subcls1_def)
lemma subcls_compareable:
@@ -134,11 +134,11 @@
\<rbrakk> \<Longrightarrow> G\<turnstile>X \<preceq>\<^sub>C Y \<or> G\<turnstile>Y \<preceq>\<^sub>C X"
by (rule triangle_lemma) (auto intro: single_inheritance)
-lemma subcls1_irrefl: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D; ws_prog G \<rbrakk>
+lemma subcls1_irrefl: "\<lbrakk>G\<turnstile>C \<prec>\<^sub>C1 D; ws_prog G \<rbrakk>
\<Longrightarrow> C \<noteq> D"
proof
assume ws: "ws_prog G" and
- subcls1: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D" and
+ subcls1: "G\<turnstile>C \<prec>\<^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\<turnstile>C \<prec>\<^sub>C\<^sub>1 D"
+ assume subcls1_C_D: "G\<turnstile>C \<prec>\<^sub>C1 D"
then obtain c where
"C\<noteq>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\<turnstile>C \<prec>\<^sub>C\<^sub>1 Z" and
+ assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C1 Z" and
subcls_Z_D: "G\<turnstile>Z \<prec>\<^sub>C D" and
nsubcls_D_Z: "\<not> G\<turnstile>D \<prec>\<^sub>C Z"
show "\<not> G\<turnstile>D \<prec>\<^sub>C C"
@@ -213,13 +213,13 @@
then show ?thesis
proof (induct rule: converse_trancl_induct)
fix C
- assume "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 D"
+ assume "G\<turnstile>C \<prec>\<^sub>C1 D"
with ws
show "C\<noteq>D"
by (blast dest: subcls1_irrefl)
next
fix C Z
- assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 Z" and
+ assume subcls1_C_Z: "G\<turnstile>C \<prec>\<^sub>C1 Z" and
subcls_Z_D: "G\<turnstile>Z \<prec>\<^sub>C D" and
neq_Z_D: "Z \<noteq> D"
show "C\<noteq>D"
@@ -298,7 +298,7 @@
assume clsC: "class G C = Some c"
assume subcls_C_C: "G\<turnstile>C \<prec>\<^sub>C D"
then obtain S where
- "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 S" and
+ "G\<turnstile>C \<prec>\<^sub>C1 S" and
subclseq_S_D: "G\<turnstile>S \<preceq>\<^sub>C D"
by (blast dest: tranclD)
with clsC
@@ -341,9 +341,9 @@
where
direct: "G\<turnstile>C\<leadsto>1J \<spacespace>\<spacespace> \<Longrightarrow> G\<turnstile>C\<leadsto>J"
| subint: "\<lbrakk>G\<turnstile>C\<leadsto>1I; G\<turnstile>I\<preceq>I J\<rbrakk> \<Longrightarrow> G\<turnstile>C\<leadsto>J"
-| subcls1: "\<lbrakk>G\<turnstile>C\<prec>\<^sub>C\<^sub>1D; G\<turnstile>D\<leadsto>J \<rbrakk> \<Longrightarrow> G\<turnstile>C\<leadsto>J"
+| subcls1: "\<lbrakk>G\<turnstile>C\<prec>\<^sub>C1D; G\<turnstile>D\<leadsto>J \<rbrakk> \<Longrightarrow> G\<turnstile>C\<leadsto>J"
-lemma implmtD: "G\<turnstile>C\<leadsto>J \<Longrightarrow> (\<exists>I. G\<turnstile>C\<leadsto>1I \<and> G\<turnstile>I\<preceq>I J) \<or> (\<exists>D. G\<turnstile>C\<prec>\<^sub>C\<^sub>1D \<and> G\<turnstile>D\<leadsto>J)"
+lemma implmtD: "G\<turnstile>C\<leadsto>J \<Longrightarrow> (\<exists>I. G\<turnstile>C\<leadsto>1I \<and> G\<turnstile>I\<preceq>I J) \<or> (\<exists>D. G\<turnstile>C\<prec>\<^sub>C1D \<and> G\<turnstile>D\<leadsto>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 \<Rightarrow> [ty list, ty list] \<Rightarrow> bool" ("_\<turnstile>_[\<preceq>]_" [71,71,71] 70)
+definition widens :: "prog \<Rightarrow> [ty list, ty list] \<Rightarrow> bool" ("_\<turnstile>_[\<preceq>]_" [71,71,71] 70) where
"G\<turnstile>Ts[\<preceq>]Ts' \<equiv> list_all2 (\<lambda>T T'. G\<turnstile>T\<preceq>T') Ts Ts'"
lemma widens_Nil [simp]: "G\<turnstile>[][\<preceq>][]"
--- a/src/HOL/Bali/TypeSafe.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/TypeSafe.thy Mon Mar 01 13:40:23 2010 +0100
@@ -95,17 +95,13 @@
section "result conformance"
-constdefs
- assign_conforms :: "st \<Rightarrow> (val \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> ty \<Rightarrow> env' \<Rightarrow> bool"
- ("_\<le>|_\<preceq>_\<Colon>\<preceq>_" [71,71,71,71] 70)
+definition assign_conforms :: "st \<Rightarrow> (val \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> ty \<Rightarrow> env' \<Rightarrow> bool" ("_\<le>|_\<preceq>_\<Colon>\<preceq>_" [71,71,71,71] 70) where
"s\<le>|f\<preceq>T\<Colon>\<preceq>E \<equiv>
(\<forall>s' w. Norm s'\<Colon>\<preceq>E \<longrightarrow> fst E,s'\<turnstile>w\<Colon>\<preceq>T \<longrightarrow> s\<le>|s' \<longrightarrow> assign f w (Norm s')\<Colon>\<preceq>E) \<and>
(\<forall>s' w. error_free s' \<longrightarrow> (error_free (assign f w s')))"
-constdefs
- rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool"
- ("_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_" [71,71,71,71,71,71] 70)
+definition rconf :: "prog \<Rightarrow> lenv \<Rightarrow> st \<Rightarrow> term \<Rightarrow> vals \<Rightarrow> tys \<Rightarrow> bool" ("_,_,_\<turnstile>_\<succ>_\<Colon>\<preceq>_" [71,71,71,71,71,71] 70) where
"G,L,s\<turnstile>t\<succ>v\<Colon>\<preceq>T
\<equiv> case T of
Inl T \<Rightarrow> if (\<exists> var. t=In2 var)
@@ -330,11 +326,8 @@
declare fun_upd_apply [simp del]
-
-constdefs
- DynT_prop::"[prog,inv_mode,qtname,ref_ty] \<Rightarrow> bool"
- ("_\<turnstile>_\<rightarrow>_\<preceq>_"[71,71,71,71]70)
- "G\<turnstile>mode\<rightarrow>D\<preceq>t \<equiv> mode = IntVir \<longrightarrow> is_class G D \<and>
+definition DynT_prop :: "[prog,inv_mode,qtname,ref_ty] \<Rightarrow> bool" ("_\<turnstile>_\<rightarrow>_\<preceq>_"[71,71,71,71]70) where
+ "G\<turnstile>mode\<rightarrow>D\<preceq>t \<equiv> mode = IntVir \<longrightarrow> is_class G D \<and>
(if (\<exists>T. t=ArrayT T) then D=Object else G\<turnstile>Class D\<preceq>RefT t)"
lemma DynT_propI:
--- a/src/HOL/Bali/WellForm.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/WellForm.thy Mon Mar 01 13:40:23 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 \<Rightarrow> pname \<Rightarrow> fdecl \<Rightarrow> bool"
+definition wf_fdecl :: "prog \<Rightarrow> pname \<Rightarrow> fdecl \<Rightarrow> bool" where
"wf_fdecl G P \<equiv> \<lambda>(fn,f). is_acc_type G P (type f)"
lemma wf_fdecl_def2: "\<And>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 \<Rightarrow> pname \<Rightarrow> sig \<Rightarrow> mhead \<Rightarrow> bool"
+definition wf_mhead :: "prog \<Rightarrow> pname \<Rightarrow> sig \<Rightarrow> mhead \<Rightarrow> bool" where
"wf_mhead G P \<equiv> \<lambda> sig mh. length (parTs sig) = length (pars mh) \<and>
\<spacespace> ( \<forall>T\<in>set (parTs sig). is_acc_type G P T) \<and>
is_acc_type G P (resTy mh) \<and>
@@ -78,7 +76,7 @@
\end{itemize}
*}
-constdefs callee_lcl:: "qtname \<Rightarrow> sig \<Rightarrow> methd \<Rightarrow> lenv"
+definition callee_lcl :: "qtname \<Rightarrow> sig \<Rightarrow> methd \<Rightarrow> lenv" where
"callee_lcl C sig m
\<equiv> \<lambda> k. (case k of
EName e
@@ -88,12 +86,11 @@
| Res \<Rightarrow> Some (resTy m))
| This \<Rightarrow> if is_static m then None else Some (Class C))"
-constdefs parameters :: "methd \<Rightarrow> lname set"
+definition parameters :: "methd \<Rightarrow> lname set" where
"parameters m \<equiv> set (map (EName \<circ> VNam) (pars m))
\<union> (if (static m) then {} else {This})"
-constdefs
- wf_mdecl :: "prog \<Rightarrow> qtname \<Rightarrow> mdecl \<Rightarrow> bool"
+definition wf_mdecl :: "prog \<Rightarrow> qtname \<Rightarrow> mdecl \<Rightarrow> bool" where
"wf_mdecl G C \<equiv>
\<lambda>(sig,m).
wf_mhead G (pid C) sig (mhead m) \<and>
@@ -219,8 +216,7 @@
superinterfaces widens to each of the corresponding result types
\end{itemize}
*}
-constdefs
- wf_idecl :: "prog \<Rightarrow> idecl \<Rightarrow> bool"
+definition wf_idecl :: "prog \<Rightarrow> idecl \<Rightarrow> bool" where
"wf_idecl G \<equiv>
\<lambda>(I,i).
ws_idecl G I (isuperIfs i) \<and>
@@ -321,8 +317,7 @@
\end{itemize}
*}
(* to Table *)
-constdefs entails:: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool"
- ("_ entails _" 20)
+definition entails :: "('a,'b) table \<Rightarrow> ('b \<Rightarrow> bool) \<Rightarrow> bool" ("_ entails _" 20) where
"t entails P \<equiv> \<forall>k. \<forall> x \<in> 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 \<Rightarrow> cdecl \<Rightarrow> bool"
+definition wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool" where
"wf_cdecl G \<equiv>
\<lambda>(C,c).
\<not>is_iface G C \<and>
@@ -361,8 +355,7 @@
))"
(*
-constdefs
- wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool"
+definition wf_cdecl :: "prog \<Rightarrow> cdecl \<Rightarrow> bool" where
"wf_cdecl G \<equiv>
\<lambda>(C,c).
\<not>is_iface G C \<and>
@@ -518,8 +511,7 @@
\item all defined classes are wellformed
\end{itemize}
*}
-constdefs
- wf_prog :: "prog \<Rightarrow> bool"
+definition wf_prog :: "prog \<Rightarrow> bool" where
"wf_prog G \<equiv> let is = ifaces G; cs = classes G in
ObjectC \<in> set cs \<and>
(\<forall> m\<in>set Object_mdecls. accmodi m \<noteq> Package) \<and>
@@ -919,7 +911,7 @@
inheritable: "G \<turnstile>Method old inheritable_in pid C" and
subclsC: "G\<turnstile>C\<prec>\<^sub>C declclass old"
from cls_C neq_C_Obj
- have super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
+ have super: "G\<turnstile>C \<prec>\<^sub>C1 super c"
by (rule subcls1I)
from wf cls_C neq_C_Obj
have accessible_super: "G\<turnstile>(Class (super c)) accessible_in (pid C)"
@@ -1385,7 +1377,7 @@
moreover note wf False cls_C
ultimately have "G\<turnstile>super c \<preceq>\<^sub>C declclass m"
by (auto intro: Hyp [rule_format])
- moreover from cls_C False have "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c" by (rule subcls1I)
+ moreover from cls_C False have "G\<turnstile>C \<prec>\<^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\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
+ have subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C1 super c"
by (rule subcls1I)
then have "G\<turnstile>C \<prec>\<^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\<turnstile>(Class (super c)) accessible_in (pid C)" and
- subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C\<^sub>1 super c"
+ subcls1_C_super: "G\<turnstile>C \<prec>\<^sub>C1 super c"
by (auto dest: wf_prog_cdecl wf_cdecl_supD is_acc_classD
intro: subcls1I)
show "\<exists>new. ?Constraint C new old"
--- a/src/HOL/Bali/WellType.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Bali/WellType.thy Mon Mar 01 13:40:23 2010 +0100
@@ -53,11 +53,10 @@
emhead = "ref_ty \<times> mhead"
--{* Some mnemotic selectors for emhead *}
-constdefs
- "declrefT" :: "emhead \<Rightarrow> ref_ty"
+definition "declrefT" :: "emhead \<Rightarrow> ref_ty" where
"declrefT \<equiv> fst"
- "mhd" :: "emhead \<Rightarrow> mhead"
+definition "mhd" :: "emhead \<Rightarrow> mhead" where
"mhd \<equiv> 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 \<equiv> \<lambda>a. None"
- invmode :: "('a::type)member_scheme \<Rightarrow> expr \<Rightarrow> inv_mode"
+definition invmode :: "('a::type)member_scheme \<Rightarrow> expr \<Rightarrow> inv_mode" where
"invmode m e \<equiv> if is_static m
then Static
else if e=Super then SuperM else IntVir"
--- a/src/HOL/Decision_Procs/Cooper.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Decision_Procs/Cooper.thy Mon Mar 01 13:40:23 2010 +0100
@@ -293,10 +293,10 @@
by (induct p, simp_all)
-constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
+definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where
"djf f p q \<equiv> (if q=T then T else if q=F then f p else
(let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
-constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
+definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
"evaldjf f ps \<equiv> 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 \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
"DJ f p \<equiv> evaldjf f (disjuncts p)"
lemma DJ: assumes fdj: "\<forall> 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<m \<or> ((n = m) \<and> lex_ns (ns,ms))) "
-constdefs lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool"
+definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" where
"lex_bnd t s \<equiv> lex_ns (bnds t, bnds s)"
consts
@@ -455,10 +455,10 @@
lemma nummul_nb: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul i t)"
by (induct t rule: nummul.induct, auto simp add: numadd_nb)
-constdefs numneg :: "num \<Rightarrow> num"
+definition numneg :: "num \<Rightarrow> num" where
"numneg t \<equiv> nummul (- 1) t"
-constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num"
+definition numsub :: "num \<Rightarrow> num \<Rightarrow> num" where
"numsub s t \<equiv> (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 \<Longrightarrow> bound0 (not p)"
by (cases p, auto)
-constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"conj p q \<equiv> (if (p = F \<or> 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 \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
@@ -515,7 +515,7 @@
lemma conj_nb: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
using conj_def by auto
-constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"disj p q \<equiv> (if (p = T \<or> 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: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
using disj_def by auto
-constdefs imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"imp p q \<equiv> (if (p = F \<or> 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 \<or> q=T",simp_all add: imp_def,cases p) (simp_all add: not)
@@ -534,7 +534,7 @@
lemma imp_nb: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
using imp_def by (cases "p=F \<or> q=T",simp_all add: imp_def,cases p) simp_all
-constdefs iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> 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 "(\<exists> x. Ifm bbs (x#bs) p) = ((\<exists> j\<in> {1 .. d}. Ifm bbs (j#bs) (minusinf p)) \<or> (\<exists> j\<in> {1.. d}. \<exists> b\<in> (Inum (i#bs)) ` set (\<beta> p). Ifm bbs ((b+j)#bs) p))"
using cp_thm[OF lp up dd dp,where i="i"] by auto
-constdefs unit:: "fm \<Rightarrow> fm \<times> num list \<times> int"
+definition unit :: "fm \<Rightarrow> fm \<times> num list \<times> int" where
"unit p \<equiv> (let p' = zlfm p ; l = \<zeta> p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\<beta> p' l); d = \<delta> q;
B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q))
in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
@@ -1814,7 +1814,7 @@
qed
(* Cooper's Algorithm *)
-constdefs cooper :: "fm \<Rightarrow> fm"
+definition cooper :: "fm \<Rightarrow> fm" where
"cooper p \<equiv>
(let (q,B,d) = unit p; js = iupt 1 d;
mq = simpfm (minusinf q);
--- a/src/HOL/Decision_Procs/Ferrack.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Decision_Procs/Ferrack.thy Mon Mar 01 13:40:23 2010 +0100
@@ -169,26 +169,26 @@
lemma not[simp]: "Ifm bs (not p) = Ifm bs (NOT p)"
by (cases p) auto
-constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"conj p q \<equiv> (if (p = F \<or> 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 \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
-constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"disj p q \<equiv> (if (p = T \<or> 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 \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
-constdefs imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"imp p q \<equiv> (if (p = F \<or> q=T \<or> 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 \<or> q=T",simp_all add: imp_def)
-constdefs iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"iff p q \<equiv> (if (p = q) then T else if (p = NOT q \<or> 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 \<Longrightarrow> qfree p"
by (induct p, simp_all)
-constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
+definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where
"djf f p q \<equiv> (if q=T then T else if q=F then f p else
(let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
-constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
+definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
"evaldjf f ps \<equiv> 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 \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
"DJ f p \<equiv> evaldjf f (disjuncts p)"
lemma DJ: assumes fdj: "\<forall> p q. Ifm bs (f (Or p q)) = Ifm bs (Or (f p) (f q))"
@@ -653,10 +653,10 @@
lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
by (induct t rule: nummul.induct, auto )
-constdefs numneg :: "num \<Rightarrow> num"
+definition numneg :: "num \<Rightarrow> num" where
"numneg t \<equiv> nummul t (- 1)"
-constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num"
+definition numsub :: "num \<Rightarrow> num \<Rightarrow> num" where
"numsub s t \<equiv> (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 \<times> int) \<Rightarrow> num \<times> int"
+definition simp_num_pair :: "(num \<times> int) \<Rightarrow> num \<times> int" where
"simp_num_pair \<equiv> (\<lambda> (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 \<Rightarrow> fm"
+definition ferrack :: "fm \<Rightarrow> fm" where
"ferrack p \<equiv> (let p' = rlfm (simpfm p); mp = minusinf p'; pp = plusinf p'
in if (mp = T \<or> pp = T) then T else
(let U = remdps(map simp_num_pair
--- a/src/HOL/Decision_Procs/MIR.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Decision_Procs/MIR.thy Mon Mar 01 13:40:23 2010 +0100
@@ -566,7 +566,7 @@
thus ?thesis by (simp only: list_all_iff)
qed
-constdefs DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
"DJ f p \<equiv> evaldjf f (disjuncts p)"
lemma DJ: assumes fdj: "\<forall> 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<m \<or> ((n = m) \<and> lex_ns (ns,ms))) "
-constdefs lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool"
+definition lex_bnd :: "num \<Rightarrow> num \<Rightarrow> bool" where
"lex_bnd t s \<equiv> lex_ns (bnds t, bnds s)"
consts
@@ -873,10 +873,10 @@
lemma nummul_nb[simp]: "\<And> i. numbound0 t \<Longrightarrow> numbound0 (nummul t i)"
by (induct t rule: nummul.induct, auto)
-constdefs numneg :: "num \<Rightarrow> num"
+definition numneg :: "num \<Rightarrow> num" where
"numneg t \<equiv> nummul t (- 1)"
-constdefs numsub :: "num \<Rightarrow> num \<Rightarrow> num"
+definition numsub :: "num \<Rightarrow> num \<Rightarrow> num" where
"numsub s t \<equiv> (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 \<times> int) \<Rightarrow> num \<times> int"
+definition simp_num_pair :: "(num \<times> int) \<Rightarrow> num \<times> int" where
"simp_num_pair \<equiv> (\<lambda> (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 \<Longrightarrow> bound0 (not p)"
by (induct p, auto)
-constdefs conj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"conj p q \<equiv> (if (p = F \<or> 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]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (conj p q)"
using conj_def by auto
-constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"disj p q \<equiv> (if (p = T \<or> 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]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (disj p q)"
using disj_def by auto
-constdefs imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"imp p q \<equiv> (if (p = F \<or> q=T \<or> 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]: "\<lbrakk>bound0 p ; bound0 q\<rbrakk> \<Longrightarrow> bound0 (imp p q)"
using imp_def by (cases "p=F \<or> q=T \<or> p=q",simp_all add: imp_def)
-constdefs iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"iff p q \<equiv> (if (p = q) then T else if (p = not q \<or> 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 \<Rightarrow> num \<Rightarrow> (int \<times> num)"
+definition simpdvd :: "int \<Rightarrow> num \<Rightarrow> (int \<times> num)" where
"simpdvd d t \<equiv>
(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 \<Rightarrow> fm"
+definition list_conj :: "fm list \<Rightarrow> fm" where
"list_conj ps \<equiv> foldr conj ps T"
lemma list_conj: "Ifm bs (list_conj ps) = (\<forall>p\<in> 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: " \<forall>p\<in> set ps. bound0 p \<Longrightarrow> bound0 (list_conj ps)"
by (induct ps, auto simp add: list_conj_def)
-constdefs CJNB:: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
"CJNB f p \<equiv> (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\<rho> p = (\<lambda> k. p)"
-constdefs \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm"
+definition \<sigma> :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> fm" where
"\<sigma> p k t \<equiv> And (Dvd k t) (\<sigma>\<rho> p (t,k))"
lemma \<sigma>\<rho>:
@@ -3517,7 +3517,7 @@
"isrlfm (Ge (CN 0 c e)) = (c>0 \<and> numbound0 e)"
"isrlfm p = (isatom p \<and> (bound0 p))"
-constdefs fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm"
+definition fp :: "fm \<Rightarrow> int \<Rightarrow> num \<Rightarrow> int \<Rightarrow> fm" where
"fp p n s j \<equiv> (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 \<Rightarrow> num \<Rightarrow> fm) \<Rightarrow> num \<Rightarrow> fm"
+definition rsplit :: "(int \<Rightarrow> num \<Rightarrow> fm) \<Rightarrow> num \<Rightarrow> fm" where
"rsplit f a \<equiv> foldr disj (map (\<lambda> (\<phi>, n, s). conj \<phi> (f n s)) (rsplit0 a)) F"
lemma foldr_disj_map: "Ifm bs (foldr disj (map f xs) F) = (\<exists> x \<in> 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 \<Rightarrow> fm"
+definition ferrack01 :: "fm \<Rightarrow> fm" where
"ferrack01 p \<equiv> (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 (\<lambda> ((t,n),(s,m)). (Add (Mul m t) (Mul n s) , 2*n*m))
@@ -5350,7 +5350,7 @@
shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. d}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> j\<in> {1.. d}. \<exists> b\<in> (Inum (real i#bs)) ` set (\<beta> p). Ifm ((b+real j)#bs) p))"
using cp_thm[OF lp up dd dp] by auto
-constdefs unit:: "fm \<Rightarrow> fm \<times> num list \<times> int"
+definition unit :: "fm \<Rightarrow> fm \<times> num list \<times> int" where
"unit p \<equiv> (let p' = zlfm p ; l = \<zeta> p' ; q = And (Dvd l (CN 0 1 (C 0))) (a\<beta> p' l); d = \<delta> q;
B = remdups (map simpnum (\<beta> q)) ; a = remdups (map simpnum (\<alpha> q))
in if length B \<le> length a then (q,B,d) else (mirror q, a,d))"
@@ -5417,7 +5417,7 @@
qed
(* Cooper's Algorithm *)
-constdefs cooper :: "fm \<Rightarrow> fm"
+definition cooper :: "fm \<Rightarrow> fm" where
"cooper p \<equiv>
(let (q,B,d) = unit p; js = iupt (1,d);
mq = simpfm (minusinf q);
@@ -5561,7 +5561,7 @@
shows "(\<exists> (x::int). Ifm (real x#bs) p) = ((\<exists> j\<in> {1 .. \<delta> p}. Ifm (real j#bs) (minusinf p)) \<or> (\<exists> (e,c) \<in> R. \<exists> j\<in> {1.. c*(\<delta> p)}. Ifm (a#bs) (\<sigma> p c (Add e (C j)))))"
using rl_thm[OF lp] \<rho>_cong[OF iszlfm_gen[OF lp, rule_format, where y="a"] R] by simp
-constdefs chooset:: "fm \<Rightarrow> fm \<times> ((num\<times>int) list) \<times> int"
+definition chooset :: "fm \<Rightarrow> fm \<times> ((num\<times>int) list) \<times> int" where
"chooset p \<equiv> (let q = zlfm p ; d = \<delta> q;
B = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<rho> q)) ;
a = remdups (map (\<lambda> (t,k). (simpnum t,k)) (\<alpha>\<rho> q))
@@ -5621,7 +5621,7 @@
ultimately show ?thes by blast
qed
-constdefs stage:: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm"
+definition stage :: "fm \<Rightarrow> int \<Rightarrow> (num \<times> int) \<Rightarrow> fm" where
"stage p d \<equiv> (\<lambda> (e,c). evaldjf (\<lambda> j. simpfm (\<sigma> p c (Add e (C j)))) (iupt (1,c*d)))"
lemma stage:
shows "Ifm bs (stage p d (e,c)) = (\<exists> j\<in>{1 .. c*d}. Ifm bs (\<sigma> 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 \<Rightarrow> fm"
+definition redlove :: "fm \<Rightarrow> fm" where
"redlove p \<equiv>
(let (q,B,d) = chooset p;
mq = simpfm (minusinf q);
--- a/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Decision_Procs/Parametric_Ferrante_Rackoff.thy Mon Mar 01 13:40:23 2010 +0100
@@ -273,10 +273,10 @@
assumes "SORT_CONSTRAINT('a::{ring_char_0,division_by_zero, field})"
shows "allpolys isnpoly t \<Longrightarrow> isnpoly c \<Longrightarrow> allpolys isnpoly (tmmul t c)" by (induct t rule: tmmul.induct, simp_all add: Let_def polymul_norm)
-constdefs tmneg :: "tm \<Rightarrow> tm"
+definition tmneg :: "tm \<Rightarrow> tm" where
"tmneg t \<equiv> tmmul t (C (- 1,1))"
-constdefs tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm"
+definition tmsub :: "tm \<Rightarrow> tm \<Rightarrow> tm" where
"tmsub s t \<equiv> (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 \<Rightarrow> fm \<Rightarrow> fm"
+definition conj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"conj p q \<equiv> (if (p = F \<or> 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 \<or> q=F",simp_all add: conj_def) (cases p,simp_all)
-constdefs disj :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition disj :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"disj p q \<equiv> (if (p = T \<or> 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 \<or> q=T",simp_all add: disj_def) (cases p,simp_all)
-constdefs imp :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition imp :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"imp p q \<equiv> (if (p = F \<or> q=T \<or> 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 \<or> q=T",simp_all add: imp_def)
-constdefs iff :: "fm \<Rightarrow> fm \<Rightarrow> fm"
+definition iff :: "fm \<Rightarrow> fm \<Rightarrow> fm" where
"iff p q \<equiv> (if (p = q) then T else if (p = NOT q \<or> 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 \<Longrightarrow> qfree p"
by (induct p, simp_all)
-constdefs djf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm"
+definition djf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a \<Rightarrow> fm \<Rightarrow> fm" where
"djf f p q \<equiv> (if q=T then T else if q=F then f p else
(let fp = f p in case fp of T \<Rightarrow> T | F \<Rightarrow> q | _ \<Rightarrow> Or (f p) q))"
-constdefs evaldjf:: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm"
+definition evaldjf :: "('a \<Rightarrow> fm) \<Rightarrow> 'a list \<Rightarrow> fm" where
"evaldjf f ps \<equiv> 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 \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition DJ :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
"DJ f p \<equiv> evaldjf f (disjuncts p)"
lemma DJ: assumes fdj: "\<forall> 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 \<Rightarrow> fm"
+definition list_conj :: "fm list \<Rightarrow> fm" where
"list_conj ps \<equiv> foldr conj ps T"
-constdefs CJNB:: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm"
+definition CJNB :: "(fm \<Rightarrow> fm) \<Rightarrow> fm \<Rightarrow> fm" where
"CJNB f p \<equiv> (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: "(\<forall> q \<in> set (conjs p). Ifm vs bs q) = Ifm vs bs p"
by (induct p rule: conjs.induct, auto)
-constdefs list_disj :: "fm list \<Rightarrow> fm"
+definition list_disj :: "fm list \<Rightarrow> fm" where
"list_disj ps \<equiv> foldr disj ps F"
lemma list_conj: "Ifm vs bs (list_conj ps) = (\<forall>p\<in> set ps. Ifm vs bs p)"
--- a/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Decision_Procs/Reflected_Multivariate_Polynomial.thy Mon Mar 01 13:40:23 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 \<Rightarrow> (poly \<times> bool)"
+definition monic :: "poly \<Rightarrow> (poly \<times> bool)" where
"monic p \<equiv> (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 \<Rightarrow> poly"
+definition shift1 :: "poly \<Rightarrow> poly" where
"shift1 p \<equiv> CN 0\<^sub>p 0 p"
consts funpow :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a"
@@ -212,7 +212,7 @@
by pat_completeness auto
-constdefs polydivide:: "poly \<Rightarrow> poly \<Rightarrow> (nat \<times> poly)"
+definition polydivide :: "poly \<Rightarrow> poly \<Rightarrow> (nat \<times> poly)" where
"polydivide s p \<equiv> polydivide_aux (head p,degree p,p,0, s)"
fun poly_deriv_aux :: "nat \<Rightarrow> poly \<Rightarrow> poly" where
@@ -262,7 +262,7 @@
lemma isnpolyh_mono: "\<lbrakk>n' \<le> n ; isnpolyh p n\<rbrakk> \<Longrightarrow> isnpolyh p n'"
by (induct p rule: isnpolyh.induct, auto)
-constdefs isnpoly:: "poly \<Rightarrow> bool"
+definition isnpoly :: "poly \<Rightarrow> bool" where
"isnpoly p \<equiv> isnpolyh p 0"
text{* polyadd preserves normal forms *}
--- a/src/HOL/Finite_Set.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Finite_Set.thy Mon Mar 01 13:40:23 2010 +0100
@@ -2528,8 +2528,7 @@
fold1Set_insertI [intro]:
"\<lbrakk> fold_graph f a A x; a \<notin> A \<rbrakk> \<Longrightarrow> 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:
--- a/src/HOL/Fun.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Fun.thy Mon Mar 01 13:40:23 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 \<longleftrightarrow> 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
--- a/src/HOL/HOL.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/HOL.thy Mon Mar 01 13:40:23 2010 +0100
@@ -1118,8 +1118,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 \<equiv> op ==>"
lemma simp_impliesI:
@@ -1392,13 +1391,23 @@
)
*}
-constdefs
- induct_forall where "induct_forall P == \<forall>x. P x"
- induct_implies where "induct_implies A B == A \<longrightarrow> B"
- induct_equal where "induct_equal x y == x = y"
- induct_conj where "induct_conj A B == A \<and> B"
- induct_true where "induct_true == True"
- induct_false where "induct_false == False"
+definition induct_forall where
+ "induct_forall P == \<forall>x. P x"
+
+definition induct_implies where
+ "induct_implies A B == A \<longrightarrow> B"
+
+definition induct_equal where
+ "induct_equal x y == x = y"
+
+definition induct_conj where
+ "induct_conj A B == A \<and> 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 (\<lambda>x. P x))"
by (unfold atomize_all induct_forall_def)
--- a/src/HOL/Hilbert_Choice.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hilbert_Choice.thy Mon Mar 01 13:40:23 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 & (\<forall>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 & (\<forall>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
--- a/src/HOL/Hoare/Arith2.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare/Arith2.thy Mon Mar 01 13:40:23 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"
--- a/src/HOL/Hoare/Heap.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare/Heap.thy Mon Mar 01 13:40:23 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 \<Rightarrow> 'a"
-primrec "addr(Ref a) = a"
+primrec addr :: "'a ref \<Rightarrow> 'a" where
+ "addr (Ref a) = a"
section "The heap"
subsection "Paths in the heap"
-consts
- Path :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool"
-primrec
-"Path h x [] y = (x = y)"
-"Path h x (a#as) y = (x = Ref a \<and> Path h (h a) as y)"
+primrec Path :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool" where
+ "Path h x [] y \<longleftrightarrow> x = y"
+| "Path h x (a#as) y \<longleftrightarrow> x = Ref a \<and> Path h (h a) as y"
lemma [iff]: "Path h Null xs y = (xs = [] \<and> y = Null)"
apply(case_tac xs)
@@ -60,8 +57,7 @@
subsection "Non-repeating paths"
-constdefs
- distPath :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool"
+definition distPath :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 'a ref \<Rightarrow> bool" where
"distPath h x as y \<equiv> Path h x as y \<and> 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 \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> bool"
+definition List :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list \<Rightarrow> 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 \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> bool"
+definition islist :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> bool" where
"islist h p == \<exists>as. List h p as"
- list :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list"
+
+definition list :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a list" where
"list h p == SOME as. List h p as"
lemma List_conv_islist_list: "List h p as = (islist h p \<and> as = list h p)"
--- a/src/HOL/Hoare/Hoare_Logic.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare/Hoare_Logic.thy Mon Mar 01 13:40:23 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 \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
+definition Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool" where
"Valid p c q == !s s'. Sem c s s' --> s : p --> s' : q"
--- a/src/HOL/Hoare/Hoare_Logic_Abort.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare/Hoare_Logic_Abort.thy Mon Mar 01 13:40:23 2010 +0100
@@ -42,7 +42,7 @@
"Sem(While b x c) s s' =
(if s = None then s' = None else \<exists>n. iter n b (Sem c) s s')"
-constdefs Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool"
+definition Valid :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a bexp \<Rightarrow> bool" where
"Valid p c q == \<forall>s s'. Sem c s s' \<longrightarrow> s : Some ` p \<longrightarrow> s' : Some ` q"
--- a/src/HOL/Hoare/Pointer_Examples.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare/Pointer_Examples.thy Mon Mar 01 13:40:23 2010 +0100
@@ -216,10 +216,10 @@
text"This is still a bit rough, especially the proof."
-constdefs
- cor :: "bool \<Rightarrow> bool \<Rightarrow> bool"
+definition cor :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
"cor P Q == if P then True else Q"
- cand :: "bool \<Rightarrow> bool \<Rightarrow> bool"
+
+definition cand :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
"cand P Q == if P then Q else False"
consts merge :: "'a list * 'a list * ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list"
@@ -481,7 +481,7 @@
subsection "Storage allocation"
-constdefs new :: "'a set \<Rightarrow> 'a"
+definition new :: "'a set \<Rightarrow> 'a" where
"new A == SOME a. a \<notin> A"
--- a/src/HOL/Hoare/Pointers0.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare/Pointers0.thy Mon Mar 01 13:40:23 2010 +0100
@@ -73,8 +73,7 @@
subsubsection "Relational abstraction"
-constdefs
- List :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> bool"
+definition List :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list \<Rightarrow> 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 \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool"
+definition islist :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> bool" where
"islist h p == \<exists>as. List h p as"
- list :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list"
+
+definition list :: "('a::ref \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a list" where
"list h p == SOME as. List h p as"
lemma List_conv_islist_list: "List h p as = (islist h p \<and> as = list h p)"
@@ -407,7 +406,7 @@
subsection "Storage allocation"
-constdefs new :: "'a set \<Rightarrow> 'a::ref"
+definition new :: "'a set \<Rightarrow> 'a::ref" where
"new A == SOME a. a \<notin> A & a \<noteq> Null"
--- a/src/HOL/Hoare/SepLogHeap.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare/SepLogHeap.thy Mon Mar 01 13:40:23 2010 +0100
@@ -41,8 +41,7 @@
subsection "Lists on the heap"
-constdefs
- List :: "heap \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool"
+definition List :: "heap \<Rightarrow> nat \<Rightarrow> nat list \<Rightarrow> bool" where
"List h x as == Path h x as 0"
lemma [simp]: "List h x [] = (x = 0)"
--- a/src/HOL/Hoare/Separation.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare/Separation.thy Mon Mar 01 13:40:23 2010 +0100
@@ -16,20 +16,19 @@
text{* The semantic definition of a few connectives: *}
-constdefs
- ortho:: "heap \<Rightarrow> heap \<Rightarrow> bool" (infix "\<bottom>" 55)
+definition ortho :: "heap \<Rightarrow> heap \<Rightarrow> bool" (infix "\<bottom>" 55) where
"h1 \<bottom> h2 == dom h1 \<inter> dom h2 = {}"
- is_empty :: "heap \<Rightarrow> bool"
+definition is_empty :: "heap \<Rightarrow> bool" where
"is_empty h == h = empty"
- singl:: "heap \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool"
+definition singl:: "heap \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" where
"singl h x y == dom h = {x} & h x = Some y"
- star:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)"
+definition star:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)" where
"star P Q == \<lambda>h. \<exists>h1 h2. h = h1++h2 \<and> h1 \<bottom> h2 \<and> P h1 \<and> Q h2"
- wand:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)"
+definition wand:: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> bool)" where
"wand P Q == \<lambda>h. \<forall>h'. h' \<bottom> h \<and> P h' \<longrightarrow> Q(h++h')"
text{*This is what assertions look like without any syntactic sugar: *}
--- a/src/HOL/Hoare_Parallel/Gar_Coll.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare_Parallel/Gar_Coll.thy Mon Mar 01 13:40:23 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 \<Rightarrow> bool"
+definition Mut_init :: "gar_coll_state \<Rightarrow> bool" where
"Mut_init \<equiv> \<guillemotleft> T \<in> Reach \<acute>E \<and> R < length \<acute>E \<and> T < length \<acute>M \<guillemotright>"
text {* \noindent For the mutator we
@@ -40,14 +39,13 @@
@{text "\<acute>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 \<equiv> .{\<acute>Mut_init \<and> \<acute>z}. \<langle>\<acute>E:=\<acute>E[R:=(fst(\<acute>E!R), T)],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
- Color_Target :: "gar_coll_state ann_com"
+definition Color_Target :: "gar_coll_state ann_com" where
"Color_Target \<equiv> .{\<acute>Mut_init \<and> \<not>\<acute>z}. \<langle>\<acute>M:=\<acute>M[T:=Black],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
- Mutator :: "gar_coll_state ann_com"
+definition Mutator :: "gar_coll_state ann_com" where
"Mutator \<equiv>
.{\<acute>Mut_init \<and> \<acute>z}.
WHILE True INV .{\<acute>Mut_init \<and> \<acute>z}.
@@ -88,22 +86,20 @@
consts M_init :: nodes
-constdefs
- Proper_M_init :: "gar_coll_state \<Rightarrow> bool"
+definition Proper_M_init :: "gar_coll_state \<Rightarrow> bool" where
"Proper_M_init \<equiv> \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>"
- Proper :: "gar_coll_state \<Rightarrow> bool"
+definition Proper :: "gar_coll_state \<Rightarrow> bool" where
"Proper \<equiv> \<guillemotleft> Proper_Roots \<acute>M \<and> Proper_Edges(\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<guillemotright>"
- Safe :: "gar_coll_state \<Rightarrow> bool"
+definition Safe :: "gar_coll_state \<Rightarrow> bool" where
"Safe \<equiv> \<guillemotleft> Reach \<acute>E \<subseteq> Blacks \<acute>M \<guillemotright>"
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 \<equiv>
.{\<acute>Proper}.
\<acute>ind:=0;;
@@ -133,13 +129,11 @@
subsubsection {* Propagating black *}
-constdefs
- PBInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition PBInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
"PBInv \<equiv> \<guillemotleft> \<lambda>ind. \<acute>obc < Blacks \<acute>M \<or> (\<forall>i <ind. \<not>BtoW (\<acute>E!i, \<acute>M) \<or>
(\<not>\<acute>z \<and> i=R \<and> (snd(\<acute>E!R)) = T \<and> (\<exists>r. ind \<le> r \<and> r < length \<acute>E \<and> BtoW(\<acute>E!r,\<acute>M))))\<guillemotright>"
-constdefs
- Propagate_Black_aux :: "gar_coll_state ann_com"
+definition Propagate_Black_aux :: "gar_coll_state ann_com" where
"Propagate_Black_aux \<equiv>
.{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
\<acute>ind:=0;;
@@ -214,14 +208,12 @@
subsubsection {* Refining propagating black *}
-constdefs
- Auxk :: "gar_coll_state \<Rightarrow> bool"
+definition Auxk :: "gar_coll_state \<Rightarrow> bool" where
"Auxk \<equiv> \<guillemotleft>\<acute>k<length \<acute>M \<and> (\<acute>M!\<acute>k\<noteq>Black \<or> \<not>BtoW(\<acute>E!\<acute>ind, \<acute>M) \<or>
\<acute>obc<Blacks \<acute>M \<or> (\<not>\<acute>z \<and> \<acute>ind=R \<and> snd(\<acute>E!R)=T
\<and> (\<exists>r. \<acute>ind<r \<and> r<length \<acute>E \<and> BtoW(\<acute>E!r, \<acute>M))))\<guillemotright>"
-constdefs
- Propagate_Black :: " gar_coll_state ann_com"
+definition Propagate_Black :: " gar_coll_state ann_com" where
"Propagate_Black \<equiv>
.{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
\<acute>ind:=0;;
@@ -328,12 +320,10 @@
subsubsection {* Counting black nodes *}
-constdefs
- CountInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition CountInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
"CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
-constdefs
- Count :: " gar_coll_state ann_com"
+definition Count :: " gar_coll_state ann_com" where
"Count \<equiv>
.{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M
\<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
@@ -398,12 +388,10 @@
Append_to_free2: "i \<notin> Reach e
\<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)"
-constdefs
- AppendInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition AppendInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
"AppendInv \<equiv> \<guillemotleft>\<lambda>ind. \<forall>i<length \<acute>M. ind\<le>i \<longrightarrow> i\<in>Reach \<acute>E \<longrightarrow> \<acute>M!i=Black\<guillemotright>"
-constdefs
- Append :: " gar_coll_state ann_com"
+definition Append :: " gar_coll_state ann_com" where
"Append \<equiv>
.{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe}.
\<acute>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 \<equiv>
.{\<acute>Proper}.
WHILE True INV .{\<acute>Proper}.
--- a/src/HOL/Hoare_Parallel/Graph.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare_Parallel/Graph.thy Mon Mar 01 13:40:23 2010 +0100
@@ -13,20 +13,19 @@
consts Roots :: "nat set"
-constdefs
- Proper_Roots :: "nodes \<Rightarrow> bool"
+definition Proper_Roots :: "nodes \<Rightarrow> bool" where
"Proper_Roots M \<equiv> Roots\<noteq>{} \<and> Roots \<subseteq> {i. i<length M}"
- Proper_Edges :: "(nodes \<times> edges) \<Rightarrow> bool"
+definition Proper_Edges :: "(nodes \<times> edges) \<Rightarrow> bool" where
"Proper_Edges \<equiv> (\<lambda>(M,E). \<forall>i<length E. fst(E!i)<length M \<and> snd(E!i)<length M)"
- BtoW :: "(edge \<times> nodes) \<Rightarrow> bool"
+definition BtoW :: "(edge \<times> nodes) \<Rightarrow> bool" where
"BtoW \<equiv> (\<lambda>(e,M). (M!fst e)=Black \<and> (M!snd e)\<noteq>Black)"
- Blacks :: "nodes \<Rightarrow> nat set"
+definition Blacks :: "nodes \<Rightarrow> nat set" where
"Blacks M \<equiv> {i. i<length M \<and> M!i=Black}"
- Reach :: "edges \<Rightarrow> nat set"
+definition Reach :: "edges \<Rightarrow> nat set" where
"Reach E \<equiv> {x. (\<exists>path. 1<length path \<and> path!(length path - 1)\<in>Roots \<and> x=path!0
\<and> (\<forall>i<length path - 1. (\<exists>j<length E. E!j=(path!(i+1), path!i))))
\<or> x\<in>Roots}"
--- a/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Mon Mar 01 13:40:23 2010 +0100
@@ -26,24 +26,23 @@
subsection {* The Mutators *}
-constdefs
- Mul_mut_init :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition Mul_mut_init :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
"Mul_mut_init \<equiv> \<guillemotleft> \<lambda>n. n=length \<acute>Muts \<and> (\<forall>i<n. R (\<acute>Muts!i)<length \<acute>E
\<and> T (\<acute>Muts!i)<length \<acute>M) \<guillemotright>"
- Mul_Redirect_Edge :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Redirect_Edge :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Redirect_Edge j n \<equiv>
.{\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)}.
\<langle>IF T(\<acute>Muts!j) \<in> Reach \<acute>E THEN
\<acute>E:= \<acute>E[R (\<acute>Muts!j):= (fst (\<acute>E!R(\<acute>Muts!j)), T (\<acute>Muts!j))] FI,,
\<acute>Muts:= \<acute>Muts[j:= (\<acute>Muts!j) \<lparr>Z:=False\<rparr>]\<rangle>"
- Mul_Color_Target :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Color_Target :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Color_Target j n \<equiv>
.{\<acute>Mul_mut_init n \<and> \<not> Z (\<acute>Muts!j)}.
\<langle>\<acute>M:=\<acute>M[T (\<acute>Muts!j):=Black],, \<acute>Muts:=\<acute>Muts[j:= (\<acute>Muts!j) \<lparr>Z:=True\<rparr>]\<rangle>"
- Mul_Mutator :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Mutator :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Mutator j n \<equiv>
.{\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)}.
WHILE True
@@ -156,28 +155,25 @@
subsection {* The Collector *}
-constdefs
- Queue :: "mul_gar_coll_state \<Rightarrow> nat"
+definition Queue :: "mul_gar_coll_state \<Rightarrow> nat" where
"Queue \<equiv> \<guillemotleft> length (filter (\<lambda>i. \<not> Z i \<and> \<acute>M!(T i) \<noteq> Black) \<acute>Muts) \<guillemotright>"
consts M_init :: nodes
-constdefs
- Proper_M_init :: "mul_gar_coll_state \<Rightarrow> bool"
+definition Proper_M_init :: "mul_gar_coll_state \<Rightarrow> bool" where
"Proper_M_init \<equiv> \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>"
- Mul_Proper :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition Mul_Proper :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
"Mul_Proper \<equiv> \<guillemotleft> \<lambda>n. Proper_Roots \<acute>M \<and> Proper_Edges (\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<and> n=length \<acute>Muts \<guillemotright>"
- Safe :: "mul_gar_coll_state \<Rightarrow> bool"
+definition Safe :: "mul_gar_coll_state \<Rightarrow> bool" where
"Safe \<equiv> \<guillemotleft> Reach \<acute>E \<subseteq> Blacks \<acute>M \<guillemotright>"
lemmas mul_collector_defs = Proper_M_init_def Mul_Proper_def Safe_def
subsubsection {* Blackening Roots *}
-constdefs
- Mul_Blacken_Roots :: "nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Blacken_Roots :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Blacken_Roots n \<equiv>
.{\<acute>Mul_Proper n}.
\<acute>ind:=0;;
@@ -208,16 +204,14 @@
subsubsection {* Propagating Black *}
-constdefs
- Mul_PBInv :: "mul_gar_coll_state \<Rightarrow> bool"
+definition Mul_PBInv :: "mul_gar_coll_state \<Rightarrow> bool" where
"Mul_PBInv \<equiv> \<guillemotleft>\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue
\<or> (\<forall>i<\<acute>ind. \<not>BtoW(\<acute>E!i,\<acute>M)) \<and> \<acute>l\<le>\<acute>Queue\<guillemotright>"
- Mul_Auxk :: "mul_gar_coll_state \<Rightarrow> bool"
+definition Mul_Auxk :: "mul_gar_coll_state \<Rightarrow> bool" where
"Mul_Auxk \<equiv> \<guillemotleft>\<acute>l<\<acute>Queue \<or> \<acute>M!\<acute>k\<noteq>Black \<or> \<not>BtoW(\<acute>E!\<acute>ind, \<acute>M) \<or> \<acute>obc\<subset>Blacks \<acute>M\<guillemotright>"
-constdefs
- Mul_Propagate_Black :: "nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Propagate_Black :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Propagate_Black n \<equiv>
.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
\<and> (\<acute>Safe \<or> \<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)}.
@@ -296,11 +290,10 @@
subsubsection {* Counting Black Nodes *}
-constdefs
- Mul_CountInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
- "Mul_CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
+definition Mul_CountInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
+ "Mul_CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
- Mul_Count :: "nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Count :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Count n \<equiv>
.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M
\<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
@@ -396,11 +389,10 @@
Append_to_free2: "i \<notin> Reach e
\<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)"
-constdefs
- Mul_AppendInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+definition Mul_AppendInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
"Mul_AppendInv \<equiv> \<guillemotleft> \<lambda>ind. (\<forall>i. ind\<le>i \<longrightarrow> i<length \<acute>M \<longrightarrow> i\<in>Reach \<acute>E \<longrightarrow> \<acute>M!i=Black)\<guillemotright>"
- Mul_Append :: "nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Append :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Append n \<equiv>
.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe}.
\<acute>ind:=0;;
@@ -438,8 +430,7 @@
subsubsection {* Collector *}
-constdefs
- Mul_Collector :: "nat \<Rightarrow> mul_gar_coll_state ann_com"
+definition Mul_Collector :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Collector n \<equiv>
.{\<acute>Mul_Proper n}.
WHILE True INV .{\<acute>Mul_Proper n}.
--- a/src/HOL/Hoare_Parallel/OG_Hoare.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy Mon Mar 01 13:40:23 2010 +0100
@@ -27,12 +27,12 @@
consts post :: "'a ann_triple_op \<Rightarrow> 'a assn"
primrec "post (c, q) = q"
-constdefs interfree_aux :: "('a ann_com_op \<times> 'a assn \<times> 'a ann_com_op) \<Rightarrow> bool"
+definition interfree_aux :: "('a ann_com_op \<times> 'a assn \<times> 'a ann_com_op) \<Rightarrow> bool" where
"interfree_aux \<equiv> \<lambda>(co, q, co'). co'= None \<or>
(\<forall>(r,a) \<in> atomics (the co'). \<parallel>= (q \<inter> r) a q \<and>
(co = None \<or> (\<forall>p \<in> assertions (the co). \<parallel>= (p \<inter> r) a p)))"
-constdefs interfree :: "(('a ann_triple_op) list) \<Rightarrow> bool"
+definition interfree :: "(('a ann_triple_op) list) \<Rightarrow> bool" where
"interfree Ts \<equiv> \<forall>i j. i < length Ts \<and> j < length Ts \<and> i \<noteq> j \<longrightarrow>
interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) "
--- a/src/HOL/Hoare_Parallel/OG_Tactics.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy Mon Mar 01 13:40:23 2010 +0100
@@ -171,8 +171,7 @@
"\<parallel>- (q \<inter> (r \<inter> b)) a q \<Longrightarrow> 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) \<Rightarrow> bool"
+definition interfree_swap :: "('a ann_triple_op * ('a ann_triple_op) list) \<Rightarrow> bool" where
"interfree_swap == \<lambda>(x, xs). \<forall>y\<in>set xs. interfree_aux (com x, post x, com y)
\<and> interfree_aux(com y, post y, com x)"
@@ -208,7 +207,7 @@
\<Longrightarrow> interfree (map (\<lambda>k. (c k, Q k)) [a..<b])"
by(force simp add: interfree_def less_diff_conv)
-constdefs map_ann_hoare :: "(('a ann_com_op * 'a assn) list) \<Rightarrow> bool " ("[\<turnstile>] _" [0] 45)
+definition map_ann_hoare :: "(('a ann_com_op * 'a assn) list) \<Rightarrow> bool " ("[\<turnstile>] _" [0] 45) where
"[\<turnstile>] Ts == (\<forall>i<length Ts. \<exists>c q. Ts!i=(Some c, q) \<and> \<turnstile> c q)"
lemma MapAnnEmpty: "[\<turnstile>] []"
--- a/src/HOL/Hoare_Parallel/OG_Tran.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Tran.thy Mon Mar 01 13:40:23 2010 +0100
@@ -7,14 +7,13 @@
'a ann_com_op = "('a ann_com) option"
'a ann_triple_op = "('a ann_com_op \<times> 'a assn)"
-consts com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op"
-primrec "com (c, q) = c"
+primrec com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op" where
+ "com (c, q) = c"
-consts post :: "'a ann_triple_op \<Rightarrow> 'a assn"
-primrec "post (c, q) = q"
+primrec post :: "'a ann_triple_op \<Rightarrow> 'a assn" where
+ "post (c, q) = q"
-constdefs
- All_None :: "'a ann_triple_op list \<Rightarrow> bool"
+definition All_None :: "'a ann_triple_op list \<Rightarrow> bool" where
"All_None Ts \<equiv> \<forall>(c, q) \<in> set Ts. c = None"
subsection {* The Transition Relation *}
@@ -88,26 +87,24 @@
subsection {* Definition of Semantics *}
-constdefs
- ann_sem :: "'a ann_com \<Rightarrow> 'a \<Rightarrow> 'a set"
+definition ann_sem :: "'a ann_com \<Rightarrow> 'a \<Rightarrow> 'a set" where
"ann_sem c \<equiv> \<lambda>s. {t. (Some c, s) -*\<rightarrow> (None, t)}"
- ann_SEM :: "'a ann_com \<Rightarrow> 'a set \<Rightarrow> 'a set"
+definition ann_SEM :: "'a ann_com \<Rightarrow> 'a set \<Rightarrow> 'a set" where
"ann_SEM c S \<equiv> \<Union>ann_sem c ` S"
- sem :: "'a com \<Rightarrow> 'a \<Rightarrow> 'a set"
+definition sem :: "'a com \<Rightarrow> 'a \<Rightarrow> 'a set" where
"sem c \<equiv> \<lambda>s. {t. \<exists>Ts. (c, s) -P*\<rightarrow> (Parallel Ts, t) \<and> All_None Ts}"
- SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set"
+definition SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set" where
"SEM c S \<equiv> \<Union>sem c ` S "
abbreviation Omega :: "'a com" ("\<Omega>" 63)
where "\<Omega> \<equiv> While UNIV UNIV (Basic id)"
-consts fwhile :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> nat \<Rightarrow> 'a com"
-primrec
- "fwhile b c 0 = \<Omega>"
- "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)"
+primrec fwhile :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> nat \<Rightarrow> 'a com" where
+ "fwhile b c 0 = \<Omega>"
+ | "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 \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>= _// _//_)" [90,55,90] 50)
+definition com_validity :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>= _// _//_)" [90,55,90] 50) where
"\<parallel>= p c q \<equiv> SEM c p \<subseteq> q"
- ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("\<Turnstile> _ _" [60,90] 45)
+definition ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("\<Turnstile> _ _" [60,90] 45) where
"\<Turnstile> c q \<equiv> ann_SEM c (pre c) \<subseteq> q"
end
\ No newline at end of file
--- a/src/HOL/Hoare_Parallel/RG_Hoare.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy Mon Mar 01 13:40:23 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 \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool"
+definition stable :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool" where
"stable \<equiv> \<lambda>f g. (\<forall>x y. x \<in> f \<longrightarrow> (x, y) \<in> g \<longrightarrow> y \<in> f)"
inductive
@@ -39,16 +38,19 @@
\<turnstile> P sat [pre', rely', guar', post'] \<rbrakk>
\<Longrightarrow> \<turnstile> P sat [pre, rely, guar, post]"
-constdefs
- Pre :: "'a rgformula \<Rightarrow> 'a set"
+definition Pre :: "'a rgformula \<Rightarrow> 'a set" where
"Pre x \<equiv> fst(snd x)"
- Post :: "'a rgformula \<Rightarrow> 'a set"
+
+definition Post :: "'a rgformula \<Rightarrow> 'a set" where
"Post x \<equiv> snd(snd(snd(snd x)))"
- Rely :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set"
+
+definition Rely :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set" where
"Rely x \<equiv> fst(snd(snd x))"
- Guar :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set"
+
+definition Guar :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set" where
"Guar x \<equiv> fst(snd(snd(snd x)))"
- Com :: "'a rgformula \<Rightarrow> 'a com"
+
+definition Com :: "'a rgformula \<Rightarrow> 'a com" where
"Com x \<equiv> 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 \<Rightarrow> 'a par_com"
+definition ParallelCom :: "('a rgformula) list \<Rightarrow> 'a par_com" where
"ParallelCom Ps \<equiv> map (Some \<circ> fst) Ps"
lemma two:
--- a/src/HOL/Hoare_Parallel/RG_Tran.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Tran.thy Mon Mar 01 13:40:23 2010 +0100
@@ -81,8 +81,7 @@
| CptnEnv: "(P, t)#xs \<in> cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> cptn"
| CptnComp: "\<lbrakk>(P,s) -c\<rightarrow> (Q,t); (Q, t)#xs \<in> cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> cptn"
-constdefs
- cp :: "('a com) option \<Rightarrow> 'a \<Rightarrow> ('a confs) set"
+definition cp :: "('a com) option \<Rightarrow> 'a \<Rightarrow> ('a confs) set" where
"cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> cptn}"
subsubsection {* Parallel computations *}
@@ -95,14 +94,12 @@
| ParCptnEnv: "(P,t)#xs \<in> par_cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> par_cptn"
| ParCptnComp: "\<lbrakk> (P,s) -pc\<rightarrow> (Q,t); (Q,t)#xs \<in> par_cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> par_cptn"
-constdefs
- par_cp :: "'a par_com \<Rightarrow> 'a \<Rightarrow> ('a par_confs) set"
+definition par_cp :: "'a par_com \<Rightarrow> 'a \<Rightarrow> ('a par_confs) set" where
"par_cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> par_cptn}"
subsection{* Modular Definition of Computation *}
-constdefs
- lift :: "'a com \<Rightarrow> 'a conf \<Rightarrow> 'a conf"
+definition lift :: "'a com \<Rightarrow> 'a conf \<Rightarrow> 'a conf" where
"lift Q \<equiv> \<lambda>(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 \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
-constdefs
- assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a confs) set"
+definition assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a confs) set" where
"assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow>
c!i -e\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"
- comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a confs) set"
+definition comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a confs) set" where
"comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow>
c!i -c\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and>
(fst (last c) = None \<longrightarrow> snd (last c) \<in> post)}"
- com_validity :: "'a com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool"
- ("\<Turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45)
+definition com_validity :: "'a com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool"
+ ("\<Turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45) where
"\<Turnstile> P sat [pre, rely, guar, post] \<equiv>
\<forall>s. cp (Some P) s \<inter> assum(pre, rely) \<subseteq> comm(guar, post)"
subsection {* Validity for Parallel Programs. *}
-constdefs
- All_None :: "(('a com) option) list \<Rightarrow> bool"
+definition All_None :: "(('a com) option) list \<Rightarrow> bool" where
"All_None xs \<equiv> \<forall>c\<in>set xs. c=None"
- par_assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a par_confs) set"
+definition par_assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a par_confs) set" where
"par_assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow>
c!i -pe\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"
- par_comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a par_confs) set"
+definition par_comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a par_confs) set" where
"par_comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow>
c!i -pc\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and>
(All_None (fst (last c)) \<longrightarrow> snd( last c) \<in> post)}"
- par_com_validity :: "'a par_com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set
-\<Rightarrow> 'a set \<Rightarrow> bool" ("\<Turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45)
+definition par_com_validity :: "'a par_com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set
+\<Rightarrow> 'a set \<Rightarrow> bool" ("\<Turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45) where
"\<Turnstile> Ps SAT [pre, rely, guar, post] \<equiv>
\<forall>s. par_cp Ps s \<inter> par_assum(pre, rely) \<subseteq> par_comm(guar, post)"
@@ -419,23 +414,22 @@
subsubsection {* Definition of the conjoin operator *}
-constdefs
- same_length :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+definition same_length :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
"same_length c clist \<equiv> (\<forall>i<length clist. length(clist!i)=length c)"
- same_state :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+definition same_state :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
"same_state c clist \<equiv> (\<forall>i <length clist. \<forall>j<length c. snd(c!j) = snd((clist!i)!j))"
- same_program :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+definition same_program :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
"same_program c clist \<equiv> (\<forall>j<length c. fst(c!j) = map (\<lambda>x. fst(nth x j)) clist)"
- compat_label :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+definition compat_label :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
"compat_label c clist \<equiv> (\<forall>j. Suc j<length c \<longrightarrow>
(c!j -pc\<rightarrow> c!Suc j \<and> (\<exists>i<length clist. (clist!i)!j -c\<rightarrow> (clist!i)! Suc j \<and>
(\<forall>l<length clist. l\<noteq>i \<longrightarrow> (clist!l)!j -e\<rightarrow> (clist!l)! Suc j))) \<or>
(c!j -pe\<rightarrow> c!Suc j \<and> (\<forall>i<length clist. (clist!i)!j -e\<rightarrow> (clist!i)! Suc j)))"
- conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" ("_ \<propto> _" [65,65] 64)
+definition conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" ("_ \<propto> _" [65,65] 64) where
"c \<propto> clist \<equiv> (same_length c clist) \<and> (same_state c clist) \<and> (same_program c clist) \<and> (compat_label c clist)"
subsubsection {* Some previous lemmas *}
--- a/src/HOL/IOA/Solve.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/IOA/Solve.thy Mon Mar 01 13:40:23 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 &
--- a/src/HOL/Import/HOL/HOL4Base.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Import/HOL/HOL4Base.thy Mon Mar 01 13:40:23 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"
--- a/src/HOL/Import/HOL/HOL4Prob.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Import/HOL/HOL4Prob.thy Mon Mar 01 13:40:23 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.
--- a/src/HOL/Import/HOL/HOL4Real.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Import/HOL/HOL4Real.thy Mon Mar 01 13:40:23 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))"
--- a/src/HOL/Import/HOL/HOL4Vec.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Import/HOL/HOL4Vec.thy Mon Mar 01 13:40:23 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)"
--- a/src/HOL/Import/HOL/HOL4Word32.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Import/HOL/HOL4Word32.thy Mon Mar 01 13:40:23 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"
--- a/src/HOL/Import/HOL4Compat.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Import/HOL4Compat.thy Mon Mar 01 13:40:23 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 \<Rightarrow> 'b,'a] \<Rightarrow> 'b"
+definition LET :: "['a \<Rightarrow> 'b,'a] \<Rightarrow> '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 \<Rightarrow> nat"
+
+definition NUMERAL_BIT1 :: "nat \<Rightarrow> nat" where
"NUMERAL_BIT1 n == n + (n + Suc 0)"
- NUMERAL_BIT2 :: "nat \<Rightarrow> nat"
+
+definition NUMERAL_BIT2 :: "nat \<Rightarrow> nat" where
"NUMERAL_BIT2 n == n + (n + Suc (Suc 0))"
- NUMERAL :: "nat \<Rightarrow> nat"
+
+definition NUMERAL :: "nat \<Rightarrow> 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 \<Rightarrow> nat"
+definition sum :: "nat list \<Rightarrow> 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]\<Rightarrow>'b,'b,'a list] \<Rightarrow> 'b"
+definition FOLDR :: "[['a,'b]\<Rightarrow>'b,'b,'a list] \<Rightarrow> '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 \<Rightarrow> ('a * 'b) list"
+definition ZIP :: "'a list * 'b list \<Rightarrow> ('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)"
--- a/src/HOL/Import/HOLLight/HOLLight.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Import/HOLLight/HOLLight.thy Mon Mar 01 13:40:23 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)
--- a/src/HOL/Import/HOLLightCompat.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Import/HOLLightCompat.thy Mon Mar 01 13:40:23 2010 +0100
@@ -30,8 +30,7 @@
by simp
qed
-constdefs
- Pred :: "nat \<Rightarrow> nat"
+definition Pred :: "nat \<Rightarrow> nat" where
"Pred n \<equiv> 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 \<Rightarrow> nat"
+definition NUMERAL_BIT0 :: "nat \<Rightarrow> nat" where
"NUMERAL_BIT0 n \<equiv> n + n"
lemma NUMERAL_BIT1_altdef: "NUMERAL_BIT1 n = Suc (n + n)"
--- a/src/HOL/Isar_Examples/Expr_Compiler.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Isar_Examples/Expr_Compiler.thy Mon Mar 01 13:40:23 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)"
--- a/src/HOL/Isar_Examples/Hoare.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Isar_Examples/Hoare.thy Mon Mar 01 13:40:23 2010 +0100
@@ -55,14 +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"
-syntax (xsymbols)
- Valid :: "'a bexp => 'a com => 'a bexp => bool"
- ("(3\<turnstile> _/ (2_)/ _)" [100, 55, 100] 50)
+notation (xsymbols) Valid ("(3\<turnstile> _/ (2_)/ _)" [100, 55, 100] 50)
lemma ValidI [intro?]:
"(!!s s'. Sem c s s' ==> s : P ==> s' : Q) ==> |- P c Q"
--- a/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Isar_Examples/Mutilated_Checkerboard.thy Mon Mar 01 13:40:23 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)}"
--- a/src/HOL/Matrix/ComputeNumeral.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Matrix/ComputeNumeral.thy Mon Mar 01 13:40:23 2010 +0100
@@ -20,7 +20,7 @@
lemmas bitiszero = iszero1 iszero2 iszero3 iszero4
(* lezero for bit strings *)
-constdefs
+definition
"lezero x == (x \<le> 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)"
--- a/src/HOL/Matrix/Matrix.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Matrix/Matrix.thy Mon Mar 01 13:40:23 2010 +0100
@@ -24,10 +24,10 @@
apply (rule Abs_matrix_induct)
by (simp add: Abs_matrix_inverse matrix_def)
-constdefs
- nrows :: "('a::zero) matrix \<Rightarrow> nat"
+definition nrows :: "('a::zero) matrix \<Rightarrow> 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 \<Rightarrow> nat"
+
+definition ncols :: "('a::zero) matrix \<Rightarrow> 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 \<Rightarrow> 'a infmatrix"
+definition transpose_infmatrix :: "'a infmatrix \<Rightarrow> 'a infmatrix" where
"transpose_infmatrix A j i == A i j"
- transpose_matrix :: "('a::zero) matrix \<Rightarrow> 'a matrix"
+
+definition transpose_matrix :: "('a::zero) matrix \<Rightarrow> '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 \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix"
+definition apply_infmatrix :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix" where
"apply_infmatrix f == % A. (% j i. f (A j i))"
- apply_matrix :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix"
+
+definition apply_matrix :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix" where
"apply_matrix f == % A. Abs_matrix (apply_infmatrix f (Rep_matrix A))"
- combine_infmatrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix \<Rightarrow> 'c infmatrix"
+
+definition combine_infmatrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a infmatrix \<Rightarrow> 'b infmatrix \<Rightarrow> 'c infmatrix" where
"combine_infmatrix f == % A B. (% j i. f (A j i) (B j i))"
- combine_matrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix \<Rightarrow> ('c::zero) matrix"
+
+definition combine_matrix :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a::zero) matrix \<Rightarrow> ('b::zero) matrix \<Rightarrow> ('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 \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool"
+definition commutative :: "('a \<Rightarrow> 'a \<Rightarrow> 'b) \<Rightarrow> bool" where
"commutative f == ! x y. f x y = f y x"
-associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
+
+definition associative :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 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 \<Longrightarrow> ncols A <= q \<Longrightarrow> ncols B <= q \<Longrightarrow> ncols(combine_matrix f A B) <= q"
by (simp add: ncols_le)
-constdefs
- zero_r_neutral :: "('a \<Rightarrow> 'b::zero \<Rightarrow> 'a) \<Rightarrow> bool"
+definition zero_r_neutral :: "('a \<Rightarrow> 'b::zero \<Rightarrow> 'a) \<Rightarrow> bool" where
"zero_r_neutral f == ! a. f a 0 = a"
- zero_l_neutral :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
+
+definition zero_l_neutral :: "('a::zero \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool" where
"zero_l_neutral f == ! a. f 0 a = a"
- zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool"
+
+definition zero_closed :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> bool" where
"zero_closed f == (!x. f x 0 = 0) & (!y. f 0 y = 0)"
consts foldseq :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a"
@@ -648,10 +651,10 @@
then show ?concl by simp
qed
-constdefs
- mult_matrix_n :: "nat \<Rightarrow> (('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix"
+definition mult_matrix_n :: "nat \<Rightarrow> (('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> '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) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> 'c matrix"
+
+definition mult_matrix :: "(('a::zero) \<Rightarrow> ('b::zero) \<Rightarrow> ('c::zero)) \<Rightarrow> ('c \<Rightarrow> 'c \<Rightarrow> 'c) \<Rightarrow> 'a matrix \<Rightarrow> 'b matrix \<Rightarrow> '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 \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool"
+definition r_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> bool" where
"r_distributive fmul fadd == ! a u v. fmul a (fadd u v) = fadd (fmul a u) (fmul a v)"
- l_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
+
+definition l_distributive :: "('a \<Rightarrow> 'b \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
"l_distributive fmul fadd == ! a u v. fmul (fadd u v) a = fadd (fmul u a) (fmul v a)"
- distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool"
+
+definition distributive :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> bool" where
"distributive fmul fadd == l_distributive fmul fadd & r_distributive fmul fadd"
lemma max1: "!! a x y. (a::nat) <= x \<Longrightarrow> 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 \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> 'a matrix"
+definition singleton_matrix :: "nat \<Rightarrow> nat \<Rightarrow> ('a::zero) \<Rightarrow> '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 \<Rightarrow> int \<Rightarrow> int \<Rightarrow> 'a matrix"
+
+definition move_matrix :: "('a::zero) matrix \<Rightarrow> int \<Rightarrow> int \<Rightarrow> '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 \<Rightarrow> nat \<Rightarrow> 'a matrix"
+
+definition take_rows :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> '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 \<Rightarrow> nat \<Rightarrow> 'a matrix"
+
+definition take_columns :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> '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 \<Rightarrow> nat \<Rightarrow> 'a matrix"
+definition column_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix" where
"column_of_matrix A n == take_columns (move_matrix A 0 (- int n)) 1"
- row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> 'a matrix"
+
+definition row_of_matrix :: "('a::zero) matrix \<Rightarrow> nat \<Rightarrow> '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 \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
+definition foldmatrix :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a" where
"foldmatrix f g A m n == foldseq_transposed g (% j. foldseq f (A j) n) m"
- foldmatrix_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 'a"
+
+definition foldmatrix_transposed :: "('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> ('a infmatrix) \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> '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 \<Rightarrow> 'a matrix \<Rightarrow> bool"
+definition right_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
"right_inverse_matrix A X == (A * X = one_matrix (max (nrows A) (ncols X))) \<and> nrows X \<le> ncols A"
- left_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
+
+definition left_inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
"left_inverse_matrix A X == (X * A = one_matrix (max(nrows X) (ncols A))) \<and> ncols X \<le> nrows A"
- inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool"
+
+definition inverse_matrix :: "('a::{ring_1}) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
"inverse_matrix A X == (right_inverse_matrix A X) \<and> (left_inverse_matrix A X)"
lemma right_inverse_matrix_dim: "right_inverse_matrix A X \<Longrightarrow> 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) \<Rightarrow> 'a matrix \<Rightarrow> 'a matrix"
+definition scalar_mult :: "('a::ring) \<Rightarrow> 'a matrix \<Rightarrow> 'a matrix" where
"scalar_mult a m == apply_matrix (op * a) m"
lemma scalar_mult_zero[simp]: "scalar_mult y 0 = 0"
--- a/src/HOL/Matrix/SparseMatrix.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Matrix/SparseMatrix.thy Mon Mar 01 13:40:23 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 \<Rightarrow> 'a matrix \<Rightarrow> bool"
+definition disj_matrices :: "('a::zero) matrix \<Rightarrow> 'a matrix \<Rightarrow> bool" where
"disj_matrices A B == (! j i. (Rep_matrix A j i \<noteq> 0) \<longrightarrow> (Rep_matrix B j i = 0)) & (! j i. (Rep_matrix B j i \<noteq> 0) \<longrightarrow> (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 \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
+definition diff_spmat :: "('a::lattice_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat" where
"diff_spmat A B == add_spmat A (minus_spmat B)"
lemma sorted_spmat_diff_spmat: "sorted_spmat A \<Longrightarrow> sorted_spmat B \<Longrightarrow> 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 \<Rightarrow> bool"
+definition sorted_sparse_matrix :: "'a spmat \<Rightarrow> bool" where
"sorted_sparse_matrix A == (sorted_spvec A) & (sorted_spmat A)"
lemma sorted_sparse_matrix_imp_spvec: "sorted_sparse_matrix A \<Longrightarrow> sorted_spvec A"
@@ -1014,8 +1011,7 @@
apply (simp_all add: sorted_nprt_spvec)
done
-constdefs
- mult_est_spmat :: "('a::lattice_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat"
+definition mult_est_spmat :: "('a::lattice_ring) spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> 'a spmat \<Rightarrow> '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))))"
--- a/src/HOL/Metis_Examples/BigO.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Metis_Examples/BigO.thy Mon Mar 01 13:40:23 2010 +0100
@@ -1099,9 +1099,7 @@
subsection {* Less than or equal to *}
-constdefs
- lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)"
- (infixl "<o" 70)
+definition lesso :: "('a => 'b::linordered_idom) => ('a => 'b) => ('a => 'b)" (infixl "<o" 70) where
"f <o g == (%x. max (f x - g x) 0)"
lemma bigo_lesseq1: "f =o O(h) ==> ALL x. abs (g x) <= abs (f x) ==>
--- a/src/HOL/Metis_Examples/Message.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Metis_Examples/Message.thy Mon Mar 01 13:40:23 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. \<exists>X. Crypt K X \<in> H}"
--- a/src/HOL/Metis_Examples/Tarski.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/Metis_Examples/Tarski.thy Mon Mar 01 13:40:23 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 == \<forall>x\<in>A. \<forall>y\<in>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 &
(\<forall>y \<in> 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 &
(\<forall>y \<in> 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. \<forall>y\<in>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. \<forall>y\<in>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 & (\<forall>y\<in>S. (y,L): order po) &
(\<forall>z\<in>pset po. (\<forall>y\<in>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 & (\<forall>y\<in>S. (G,y): order po) &
(\<forall>z \<in> pset po. (\<forall>y\<in>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 &
(\<forall>S. S \<subseteq> pset cl --> (\<exists>L. isLub S cl L)) &
(\<forall>S. S \<subseteq> pset cl --> (\<exists>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 \<subseteq> pset cl &
@@ -82,8 +79,7 @@
sublattice_syntax :: "['a set, 'a potype] => bool" ("_ <<= _" [51, 50] 50)
where "S <<= cl \<equiv> 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 =
--- a/src/HOL/MicroJava/BV/Altern.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/MicroJava/BV/Altern.thy Mon Mar 01 13:40:23 2010 +0100
@@ -8,19 +8,18 @@
imports BVSpec
begin
-constdefs
- check_type :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state \<Rightarrow> bool"
+definition check_type :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state \<Rightarrow> bool" where
"check_type G mxs mxr s \<equiv> s \<in> states G mxs mxr"
- wt_instr_altern :: "[instr,jvm_prog,ty,method_type,nat,nat,p_count,
- exception_table,p_count] \<Rightarrow> bool"
+definition wt_instr_altern :: "[instr,jvm_prog,ty,method_type,nat,nat,p_count,
+ exception_table,p_count] \<Rightarrow> bool" where
"wt_instr_altern i G rT phi mxs mxr max_pc et pc \<equiv>
app i G mxs rT pc et (phi!pc) \<and>
check_type G mxs mxr (OK (phi!pc)) \<and>
(\<forall>(pc',s') \<in> set (eff i G pc et (phi!pc)). pc' < max_pc \<and> G \<turnstile> s' <=' phi!pc')"
- wt_method_altern :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list,
- exception_table,method_type] \<Rightarrow> bool"
+definition wt_method_altern :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list,
+ exception_table,method_type] \<Rightarrow> bool" where
"wt_method_altern G C pTs rT mxs mxl ins et phi \<equiv>
let max_pc = length ins in
0 < max_pc \<and>
--- a/src/HOL/MicroJava/BV/BVExample.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/MicroJava/BV/BVExample.thy Mon Mar 01 13:40:23 2010 +0100
@@ -167,8 +167,7 @@
@{prop [display] "P n"}
*}
-constdefs
- intervall :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" ("_ \<in> [_, _')")
+definition intervall :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" ("_ \<in> [_, _')") where
"x \<in> [a, b) \<equiv> a \<le> x \<and> x < b"
lemma pc_0: "x < n \<Longrightarrow> (x \<in> [0, n) \<Longrightarrow> P x) \<Longrightarrow> 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 ("\<phi>\<^sub>a")
+definition phi_append :: method_type ("\<phi>\<^sub>a") where
"\<phi>\<^sub>a \<equiv> map (\<lambda>(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 ("\<phi>\<^sub>m")
+definition phi_makelist :: method_type ("\<phi>\<^sub>m") where
"\<phi>\<^sub>m \<equiv> map (\<lambda>(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 ("\<Phi>")
+definition Phi :: prog_type ("\<Phi>") where
"\<Phi> C sg \<equiv> if C = test_name \<and> sg = (makelist_name, []) then \<phi>\<^sub>m else
if C = list_name \<and> sg = (append_name, [Class list_name]) then \<phi>\<^sub>a else []"
--- a/src/HOL/MicroJava/BV/Correct.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/MicroJava/BV/Correct.thy Mon Mar 01 13:40:23 2010 +0100
@@ -9,47 +9,40 @@
imports BVSpec "../JVM/JVMExec"
begin
-constdefs
- approx_val :: "[jvm_prog,aheap,val,ty err] \<Rightarrow> bool"
+definition approx_val :: "[jvm_prog,aheap,val,ty err] \<Rightarrow> bool" where
"approx_val G h v any == case any of Err \<Rightarrow> True | OK T \<Rightarrow> G,h\<turnstile>v::\<preceq>T"
- approx_loc :: "[jvm_prog,aheap,val list,locvars_type] \<Rightarrow> bool"
+definition approx_loc :: "[jvm_prog,aheap,val list,locvars_type] \<Rightarrow> bool" where
"approx_loc G hp loc LT == list_all2 (approx_val G hp) loc LT"
- approx_stk :: "[jvm_prog,aheap,opstack,opstack_type] \<Rightarrow> bool"
+definition approx_stk :: "[jvm_prog,aheap,opstack,opstack_type] \<Rightarrow> 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] \<Rightarrow> frame \<Rightarrow> bool"
+definition correct_frame :: "[jvm_prog,aheap,state_type,nat,bytecode] \<Rightarrow> frame \<Rightarrow> bool" where
"correct_frame G hp == \<lambda>(ST,LT) maxl ins (stk,loc,C,sig,pc).
approx_stk G hp stk ST \<and> approx_loc G hp loc LT \<and>
pc < length ins \<and> length loc=length(snd sig)+maxl+1"
-
-consts
- correct_frames :: "[jvm_prog,aheap,prog_type,ty,sig,frame list] \<Rightarrow> bool"
-primrec
-"correct_frames G hp phi rT0 sig0 [] = True"
+primrec correct_frames :: "[jvm_prog,aheap,prog_type,ty,sig,frame list] \<Rightarrow> 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
+ (\<exists>ST LT rT maxs maxl ins et.
+ phi C sig ! pc = Some (ST,LT) \<and> is_class G C \<and>
+ method (G,C) sig = Some(C,rT,(maxs,maxl,ins,et)) \<and>
+ (\<exists>C' mn pTs. ins!pc = (Invoke C' mn pTs) \<and>
+ (mn,pTs) = sig0 \<and>
+ (\<exists>apTs D ST' LT'.
+ (phi C sig)!pc = Some ((rev apTs) @ (Class D) # ST', LT') \<and>
+ length apTs = length pTs \<and>
+ (\<exists>D' rT' maxs' maxl' ins' et'.
+ method (G,D) sig0 = Some(D',rT',(maxs',maxl',ins',et')) \<and>
+ G \<turnstile> rT0 \<preceq> rT') \<and>
+ correct_frame G hp (ST, LT) maxl ins f \<and>
+ 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
- (\<exists>ST LT rT maxs maxl ins et.
- phi C sig ! pc = Some (ST,LT) \<and> is_class G C \<and>
- method (G,C) sig = Some(C,rT,(maxs,maxl,ins,et)) \<and>
- (\<exists>C' mn pTs. ins!pc = (Invoke C' mn pTs) \<and>
- (mn,pTs) = sig0 \<and>
- (\<exists>apTs D ST' LT'.
- (phi C sig)!pc = Some ((rev apTs) @ (Class D) # ST', LT') \<and>
- length apTs = length pTs \<and>
- (\<exists>D' rT' maxs' maxl' ins' et'.
- method (G,D) sig0 = Some(D',rT',(maxs',maxl',ins',et')) \<and>
- G \<turnstile> rT0 \<preceq> rT') \<and>
- correct_frame G hp (ST, LT) maxl ins f \<and>
- correct_frames G hp phi rT sig frs))))"
-
-
-constdefs
- correct_state :: "[jvm_prog,prog_type,jvm_state] \<Rightarrow> bool"
- ("_,_ |-JVM _ [ok]" [51,51] 50)
+definition correct_state :: "[jvm_prog,prog_type,jvm_state] \<Rightarrow> bool"
+ ("_,_ |-JVM _ [ok]" [51,51] 50) where
"correct_state G phi == \<lambda>(xp,hp,frs).
case xp of
None \<Rightarrow> (case frs of
--- a/src/HOL/MicroJava/BV/Effect.thy Wed Feb 24 11:55:52 2010 +0100
+++ b/src/HOL/MicroJava/BV/Effect.thy Mon Mar 01 13:40:23 2010 +0100
@@ -13,27 +13,25 @@
succ_type = "(p_count \<times> state_type option) list"
text {* Program counter of successor instructions: *}
-consts
- succs :: "instr \<Rightarrow> p_count \<Rightarrow> p_count list"
-primrec
+primrec succs :: "instr \<Rightarrow> p_count \<Rightarrow> 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 \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list"
-primrec
+primrec match_any :: "jvm_prog \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> 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 \<and> pc < end_pc then catch_type#es' else es')"
-consts
- match :: "jvm_prog \<Rightarrow> xcpt \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> cname list"
-primrec
+primrec match :: "jvm_prog \<Rightarrow> xcpt \<Rightarrow> p_count \<Rightarrow> exception_table \<Rightarrow> 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 \<Rightarrow> jvm_prog \<Rightarrow> p_count \<Rightarrow> state