--- a/src/CCL/Gfp.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/CCL/Gfp.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,7 +11,7 @@
begin
definition
- gfp :: "['a set=>'a set] => 'a set" (*greatest fixed point*)
+ gfp :: "['a set=>'a set] => 'a set" where -- "greatest fixed point"
"gfp(f) == Union({u. u <= f(u)})"
(* gfp(f) is the least upper bound of {u. u <= f(u)} *)
--- a/src/CCL/Lfp.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/CCL/Lfp.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,7 +11,7 @@
begin
definition
- lfp :: "['a set=>'a set] => 'a set" (*least fixed point*)
+ lfp :: "['a set=>'a set] => 'a set" where -- "least fixed point"
"lfp(f) == Inter({u. f(u) <= u})"
(* lfp(f) is the greatest lower bound of {u. f(u) <= u} *)
--- a/src/CTT/Arith.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/CTT/Arith.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,22 +13,27 @@
subsection {* Arithmetic operators and their definitions *}
definition
- add :: "[i,i]=>i" (infixr "#+" 65)
+ add :: "[i,i]=>i" (infixr "#+" 65) where
"a#+b == rec(a, b, %u v. succ(v))"
- diff :: "[i,i]=>i" (infixr "-" 65)
+definition
+ diff :: "[i,i]=>i" (infixr "-" 65) where
"a-b == rec(b, a, %u v. rec(v, 0, %x y. x))"
- absdiff :: "[i,i]=>i" (infixr "|-|" 65)
+definition
+ absdiff :: "[i,i]=>i" (infixr "|-|" 65) where
"a|-|b == (a-b) #+ (b-a)"
- mult :: "[i,i]=>i" (infixr "#*" 70)
+definition
+ mult :: "[i,i]=>i" (infixr "#*" 70) where
"a#*b == rec(a, 0, %u v. b #+ v)"
- mod :: "[i,i]=>i" (infixr "mod" 70)
+definition
+ mod :: "[i,i]=>i" (infixr "mod" 70) where
"a mod b == rec(a, 0, %u v. rec(succ(v) |-| b, 0, %x y. succ(v)))"
- div :: "[i,i]=>i" (infixr "div" 70)
+definition
+ div :: "[i,i]=>i" (infixr "div" 70) where
"a div b == rec(a, 0, %u v. rec(succ(u) mod b, succ(v), %x y. v))"
--- a/src/CTT/Bool.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/CTT/Bool.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,16 +11,19 @@
begin
definition
- Bool :: "t"
+ Bool :: "t" where
"Bool == T+T"
- true :: "i"
+definition
+ true :: "i" where
"true == inl(tt)"
- false :: "i"
+definition
+ false :: "i" where
"false == inr(tt)"
- cond :: "[i,i,i]=>i"
+definition
+ cond :: "[i,i,i]=>i" where
"cond(a,b,c) == when(a, %u. b, %u. c)"
lemmas bool_defs = Bool_def true_def false_def cond_def
--- a/src/CTT/CTT.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/CTT/CTT.thy Fri Nov 17 02:20:03 2006 +0100
@@ -65,20 +65,21 @@
"SUM x:A. B" == "Sum(A, %x. B)"
abbreviation
- Arrow :: "[t,t]=>t" (infixr "-->" 30)
+ Arrow :: "[t,t]=>t" (infixr "-->" 30) where
"A --> B == PROD _:A. B"
- Times :: "[t,t]=>t" (infixr "*" 50)
+abbreviation
+ Times :: "[t,t]=>t" (infixr "*" 50) where
"A * B == SUM _:A. B"
notation (xsymbols)
- Elem ("(_ /\<in> _)" [10,10] 5)
- Eqelem ("(2_ =/ _ \<in>/ _)" [10,10,10] 5)
- Arrow (infixr "\<longrightarrow>" 30)
+ Elem ("(_ /\<in> _)" [10,10] 5) and
+ Eqelem ("(2_ =/ _ \<in>/ _)" [10,10,10] 5) and
+ Arrow (infixr "\<longrightarrow>" 30) and
Times (infixr "\<times>" 50)
notation (HTML output)
- Elem ("(_ /\<in> _)" [10,10] 5)
- Eqelem ("(2_ =/ _ \<in>/ _)" [10,10,10] 5)
+ Elem ("(_ /\<in> _)" [10,10] 5) and
+ Eqelem ("(2_ =/ _ \<in>/ _)" [10,10,10] 5) and
Times (infixr "\<times>" 50)
syntax (xsymbols)
--- a/src/FOL/IFOL.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/FOL/IFOL.thy Fri Nov 17 02:20:03 2006 +0100
@@ -45,7 +45,7 @@
abbreviation
- not_equal :: "['a, 'a] => o" (infixl "~=" 50)
+ not_equal :: "['a, 'a] => o" (infixl "~=" 50) where
"x ~= y == ~ (x = y)"
notation (xsymbols)
--- a/src/FOL/ex/NatClass.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/FOL/ex/NatClass.thy Fri Nov 17 02:20:03 2006 +0100
@@ -30,7 +30,7 @@
rec_Suc: "rec(Suc(m), a, f) = f(m, rec(m, a, f))"
definition
- add :: "['a::nat, 'a] => 'a" (infixl "+" 60)
+ add :: "['a::nat, 'a] => 'a" (infixl "+" 60) where
"m + n = rec(m, n, %x y. Suc(y))"
lemma Suc_n_not_n: "Suc(k) ~= (k::'a::nat)"
--- a/src/HOL/Accessible_Part.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Accessible_Part.thy Fri Nov 17 02:20:03 2006 +0100
@@ -24,7 +24,7 @@
accI: "(!!y. (y, x) \<in> r ==> y \<in> acc r) ==> x \<in> acc r"
abbreviation
- termi :: "('a \<times> 'a) set => 'a set"
+ termi :: "('a \<times> 'a) set => 'a set" where
"termi r == acc (r\<inverse>)"
--- a/src/HOL/Algebra/Coset.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Algebra/Coset.thy Fri Nov 17 02:20:03 2006 +0100
@@ -30,7 +30,7 @@
assumes coset_eq: "(\<forall>x \<in> carrier G. H #> x = x <# H)"
abbreviation
- normal_rel :: "['a set, ('a, 'b) monoid_scheme] \<Rightarrow> bool" (infixl "\<lhd>" 60)
+ normal_rel :: "['a set, ('a, 'b) monoid_scheme] \<Rightarrow> bool" (infixl "\<lhd>" 60) where
"H \<lhd> G \<equiv> normal H G"
--- a/src/HOL/Algebra/Lattice.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Algebra/Lattice.thy Fri Nov 17 02:20:03 2006 +0100
@@ -29,7 +29,7 @@
definition command cannot specialise the types. *}
definition (in order_syntax)
- less (infixl "\<sqsubset>" 50) "x \<sqsubset> y == x \<sqsubseteq> y & x ~= y"
+ less (infixl "\<sqsubset>" 50) where "x \<sqsubset> y == x \<sqsubseteq> y & x ~= y"
text {* Upper and lower bounds of a set. *}
@@ -38,35 +38,35 @@
"Upper A == {u. (ALL x. x \<in> A \<inter> L --> x \<sqsubseteq> u)} \<inter> L"
definition (in order_syntax)
- Lower :: "'a set => 'a set"
+ Lower :: "'a set => 'a set" where
"Lower A == {l. (ALL x. x \<in> A \<inter> L --> l \<sqsubseteq> x)} \<inter> L"
text {* Least and greatest, as predicate. *}
definition (in order_syntax)
- least :: "['a, 'a set] => bool"
+ least :: "['a, 'a set] => bool" where
"least l A == A \<subseteq> L & l \<in> A & (ALL x : A. l \<sqsubseteq> x)"
definition (in order_syntax)
- greatest :: "['a, 'a set] => bool"
+ greatest :: "['a, 'a set] => bool" where
"greatest g A == A \<subseteq> L & g \<in> A & (ALL x : A. x \<sqsubseteq> g)"
text {* Supremum and infimum *}
definition (in order_syntax)
- sup :: "'a set => 'a" ("\<Squnion>_" [90] 90)
+ sup :: "'a set => 'a" ("\<Squnion>_" [90] 90) where
"\<Squnion>A == THE x. least x (Upper A)"
definition (in order_syntax)
- inf :: "'a set => 'a" ("\<Sqinter>_" [90] 90)
+ inf :: "'a set => 'a" ("\<Sqinter>_" [90] 90) where
"\<Sqinter>A == THE x. greatest x (Lower A)"
definition (in order_syntax)
- join :: "['a, 'a] => 'a" (infixl "\<squnion>" 65)
+ join :: "['a, 'a] => 'a" (infixl "\<squnion>" 65) where
"x \<squnion> y == sup {x, y}"
definition (in order_syntax)
- meet :: "['a, 'a] => 'a" (infixl "\<sqinter>" 70)
+ meet :: "['a, 'a] => 'a" (infixl "\<sqinter>" 70) where
"x \<sqinter> y == inf {x, y}"
locale partial_order = order_syntax +
@@ -79,7 +79,7 @@
x \<in> L; y \<in> L; z \<in> L |] ==> x \<sqsubseteq> z"
abbreviation (in partial_order)
- less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
+ less (infixl "\<sqsubset>" 50) where "less == order_syntax.less le"
abbreviation (in partial_order)
Upper where "Upper == order_syntax.Upper L le"
abbreviation (in partial_order)
@@ -89,13 +89,13 @@
abbreviation (in partial_order)
greatest where "greatest == order_syntax.greatest L le"
abbreviation (in partial_order)
- sup ("\<Squnion>_" [90] 90) "sup == order_syntax.sup L le"
+ sup ("\<Squnion>_" [90] 90) where "sup == order_syntax.sup L le"
abbreviation (in partial_order)
- inf ("\<Sqinter>_" [90] 90) "inf == order_syntax.inf L le"
+ inf ("\<Sqinter>_" [90] 90) where "inf == order_syntax.inf L le"
abbreviation (in partial_order)
- join (infixl "\<squnion>" 65) "join == order_syntax.join L le"
+ join (infixl "\<squnion>" 65) where "join == order_syntax.join L le"
abbreviation (in partial_order)
- meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet L le"
+ meet (infixl "\<sqinter>" 70) where "meet == order_syntax.meet L le"
subsubsection {* Upper *}
@@ -207,7 +207,7 @@
"[| x \<in> L; y \<in> L |] ==> EX s. order_syntax.greatest L le s (order_syntax.Lower L le {x, y})"
abbreviation (in lattice)
- less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
+ less (infixl "\<sqsubset>" 50) where "less == order_syntax.less le"
abbreviation (in lattice)
Upper where "Upper == order_syntax.Upper L le"
abbreviation (in lattice)
@@ -217,13 +217,13 @@
abbreviation (in lattice)
greatest where "greatest == order_syntax.greatest L le"
abbreviation (in lattice)
- sup ("\<Squnion>_" [90] 90) "sup == order_syntax.sup L le"
+ sup ("\<Squnion>_" [90] 90) where "sup == order_syntax.sup L le"
abbreviation (in lattice)
- inf ("\<Sqinter>_" [90] 90) "inf == order_syntax.inf L le"
+ inf ("\<Sqinter>_" [90] 90) where "inf == order_syntax.inf L le"
abbreviation (in lattice)
- join (infixl "\<squnion>" 65) "join == order_syntax.join L le"
+ join (infixl "\<squnion>" 65) where "join == order_syntax.join L le"
abbreviation (in lattice)
- meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet L le"
+ meet (infixl "\<sqinter>" 70) where "meet == order_syntax.meet L le"
lemma (in order_syntax) least_Upper_above:
"[| least s (Upper A); x \<in> A; A \<subseteq> L |] ==> x \<sqsubseteq> s"
@@ -690,7 +690,7 @@
assumes total: "[| x \<in> L; y \<in> L |] ==> x \<sqsubseteq> y | y \<sqsubseteq> x"
abbreviation (in total_order)
- less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
+ less (infixl "\<sqsubset>" 50) where "less == order_syntax.less le"
abbreviation (in total_order)
Upper where "Upper == order_syntax.Upper L le"
abbreviation (in total_order)
@@ -700,13 +700,13 @@
abbreviation (in total_order)
greatest where "greatest == order_syntax.greatest L le"
abbreviation (in total_order)
- sup ("\<Squnion>_" [90] 90) "sup == order_syntax.sup L le"
+ sup ("\<Squnion>_" [90] 90) where "sup == order_syntax.sup L le"
abbreviation (in total_order)
- inf ("\<Sqinter>_" [90] 90) "inf == order_syntax.inf L le"
+ inf ("\<Sqinter>_" [90] 90) where "inf == order_syntax.inf L le"
abbreviation (in total_order)
- join (infixl "\<squnion>" 65) "join == order_syntax.join L le"
+ join (infixl "\<squnion>" 65) where "join == order_syntax.join L le"
abbreviation (in total_order)
- meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet L le"
+ meet (infixl "\<sqinter>" 70) where "meet == order_syntax.meet L le"
text {* Introduction rule: the usual definition of total order *}
@@ -768,7 +768,7 @@
"[| A \<subseteq> L |] ==> EX i. order_syntax.greatest L le i (order_syntax.Lower L le A)"
abbreviation (in complete_lattice)
- less (infixl "\<sqsubset>" 50) "less == order_syntax.less le"
+ less (infixl "\<sqsubset>" 50) where "less == order_syntax.less le"
abbreviation (in complete_lattice)
Upper where "Upper == order_syntax.Upper L le"
abbreviation (in complete_lattice)
@@ -778,13 +778,13 @@
abbreviation (in complete_lattice)
greatest where "greatest == order_syntax.greatest L le"
abbreviation (in complete_lattice)
- sup ("\<Squnion>_" [90] 90) "sup == order_syntax.sup L le"
+ sup ("\<Squnion>_" [90] 90) where "sup == order_syntax.sup L le"
abbreviation (in complete_lattice)
- inf ("\<Sqinter>_" [90] 90) "inf == order_syntax.inf L le"
+ inf ("\<Sqinter>_" [90] 90) where "inf == order_syntax.inf L le"
abbreviation (in complete_lattice)
- join (infixl "\<squnion>" 65) "join == order_syntax.join L le"
+ join (infixl "\<squnion>" 65) where "join == order_syntax.join L le"
abbreviation (in complete_lattice)
- meet (infixl "\<sqinter>" 70) "meet == order_syntax.meet L le"
+ meet (infixl "\<sqinter>" 70) where "meet == order_syntax.meet L le"
text {* Introduction rule: the usual definition of complete lattice *}
@@ -800,29 +800,29 @@
qed (assumption | rule complete_lattice_axioms.intro)+
definition (in order_syntax)
- top ("\<top>")
+ top ("\<top>") where
"\<top> == sup L"
definition (in order_syntax)
- bottom ("\<bottom>")
+ bottom ("\<bottom>") where
"\<bottom> == inf L"
abbreviation (in partial_order)
- top ("\<top>") "top == order_syntax.top L le"
+ top ("\<top>") where "top == order_syntax.top L le"
abbreviation (in partial_order)
- bottom ("\<bottom>") "bottom == order_syntax.bottom L le"
+ bottom ("\<bottom>") where "bottom == order_syntax.bottom L le"
abbreviation (in lattice)
- top ("\<top>") "top == order_syntax.top L le"
+ top ("\<top>") where "top == order_syntax.top L le"
abbreviation (in lattice)
- bottom ("\<bottom>") "bottom == order_syntax.bottom L le"
+ bottom ("\<bottom>") where "bottom == order_syntax.bottom L le"
abbreviation (in total_order)
- top ("\<top>") "top == order_syntax.top L le"
+ top ("\<top>") where "top == order_syntax.top L le"
abbreviation (in total_order)
- bottom ("\<bottom>") "bottom == order_syntax.bottom L le"
+ bottom ("\<bottom>") where "bottom == order_syntax.bottom L le"
abbreviation (in complete_lattice)
- top ("\<top>") "top == order_syntax.top L le"
+ top ("\<top>") where "top == order_syntax.top L le"
abbreviation (in complete_lattice)
- bottom ("\<bottom>") "bottom == order_syntax.bottom L le"
+ bottom ("\<bottom>") where "bottom == order_syntax.bottom L le"
lemma (in complete_lattice) supI:
--- a/src/HOL/Auth/CertifiedEmail.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/CertifiedEmail.thy Fri Nov 17 02:20:03 2006 +0100
@@ -8,10 +8,11 @@
theory CertifiedEmail imports Public begin
abbreviation
- TTP :: agent
+ TTP :: agent where
"TTP == Server"
- RPwd :: "agent => key"
+abbreviation
+ RPwd :: "agent => key" where
"RPwd == shrK"
--- a/src/HOL/Auth/Event.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Event.thy Fri Nov 17 02:20:03 2006 +0100
@@ -29,7 +29,7 @@
text{*The constant "spies" is retained for compatibility's sake*}
abbreviation (input)
- spies :: "event list => msg set"
+ spies :: "event list => msg set" where
"spies == knows Spy"
text{*Spy has access to his own key for spoof messages, but Server is secure*}
--- a/src/HOL/Auth/Guard/Extensions.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Guard/Extensions.thy Fri Nov 17 02:20:03 2006 +0100
@@ -86,7 +86,7 @@
by simp
abbreviation
- not_MPair :: "msg => bool"
+ not_MPair :: "msg => bool" where
"not_MPair X == ~ is_MPair X"
lemma is_MPairE: "[| is_MPair X ==> P; not_MPair X ==> P |] ==> P"
@@ -370,7 +370,7 @@
))"
abbreviation
- spies' :: "event list => msg set"
+ spies' :: "event list => msg set" where
"spies' == knows' Spy"
subsubsection{*decomposition of knows into knows' and initState*}
@@ -452,7 +452,7 @@
"knows_max A evs == knows_max' A evs Un initState A"
abbreviation
- spies_max :: "event list => msg set"
+ spies_max :: "event list => msg set" where
"spies_max evs == knows_max Spy evs"
subsubsection{*basic facts about @{term knows_max}*}
--- a/src/HOL/Auth/Guard/Guard.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Guard/Guard.thy Fri Nov 17 02:20:03 2006 +0100
@@ -163,7 +163,7 @@
subsection{*set obtained by decrypting a message*}
abbreviation (input)
- decrypt :: "msg set => key => msg => msg set"
+ decrypt :: "msg set => key => msg => msg set" where
"decrypt H K Y == insert Y (H - {Crypt K Y})"
lemma analz_decrypt: "[| Crypt K Y:H; Key (invKey K):H; Nonce n:analz H |]
--- a/src/HOL/Auth/Guard/GuardK.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Guard/GuardK.thy Fri Nov 17 02:20:03 2006 +0100
@@ -159,7 +159,7 @@
subsection{*set obtained by decrypting a message*}
abbreviation (input)
- decrypt :: "msg set => key => msg => msg set"
+ decrypt :: "msg set => key => msg => msg set" where
"decrypt H K Y == insert Y (H - {Crypt K Y})"
lemma analz_decrypt: "[| Crypt K Y:H; Key (invKey K):H; Key n:analz H |]
--- a/src/HOL/Auth/Guard/Guard_NS_Public.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Guard/Guard_NS_Public.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,19 +18,23 @@
subsection{*messages used in the protocol*}
abbreviation (input)
- ns1 :: "agent => agent => nat => event"
+ ns1 :: "agent => agent => nat => event" where
"ns1 A B NA == Says A B (Crypt (pubK B) {|Nonce NA, Agent A|})"
- ns1' :: "agent => agent => agent => nat => event"
+abbreviation (input)
+ ns1' :: "agent => agent => agent => nat => event" where
"ns1' A' A B NA == Says A' B (Crypt (pubK B) {|Nonce NA, Agent A|})"
- ns2 :: "agent => agent => nat => nat => event"
+abbreviation (input)
+ ns2 :: "agent => agent => nat => nat => event" where
"ns2 B A NA NB == Says B A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|})"
- ns2' :: "agent => agent => agent => nat => nat => event"
+abbreviation (input)
+ ns2' :: "agent => agent => agent => nat => nat => event" where
"ns2' B' B A NA NB == Says B' A (Crypt (pubK A) {|Nonce NA, Nonce NB, Agent B|})"
- ns3 :: "agent => agent => nat => event"
+abbreviation (input)
+ ns3 :: "agent => agent => nat => event" where
"ns3 A B NB == Says A B (Crypt (pubK B) (Nonce NB))"
--- a/src/HOL/Auth/Guard/Guard_OtwayRees.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Guard/Guard_OtwayRees.thy Fri Nov 17 02:20:03 2006 +0100
@@ -16,40 +16,48 @@
subsection{*messages used in the protocol*}
abbreviation
- nil :: "msg"
+ nil :: "msg" where
"nil == Number 0"
- or1 :: "agent => agent => nat => event"
+abbreviation
+ or1 :: "agent => agent => nat => event" where
"or1 A B NA ==
Says A B {|Nonce NA, Agent A, Agent B, Ciph A {|Nonce NA, Agent A, Agent B|}|}"
- or1' :: "agent => agent => agent => nat => msg => event"
+abbreviation
+ or1' :: "agent => agent => agent => nat => msg => event" where
"or1' A' A B NA X == Says A' B {|Nonce NA, Agent A, Agent B, X|}"
- or2 :: "agent => agent => nat => nat => msg => event"
+abbreviation
+ or2 :: "agent => agent => nat => nat => msg => event" where
"or2 A B NA NB X ==
Says B Server {|Nonce NA, Agent A, Agent B, X,
Ciph B {|Nonce NA, Nonce NB, Agent A, Agent B|}|}"
- or2' :: "agent => agent => agent => nat => nat => event"
+abbreviation
+ or2' :: "agent => agent => agent => nat => nat => event" where
"or2' B' A B NA NB ==
Says B' Server {|Nonce NA, Agent A, Agent B,
Ciph A {|Nonce NA, Agent A, Agent B|},
Ciph B {|Nonce NA, Nonce NB, Agent A, Agent B|}|}"
- or3 :: "agent => agent => nat => nat => key => event"
+abbreviation
+ or3 :: "agent => agent => nat => nat => key => event" where
"or3 A B NA NB K ==
Says Server B {|Nonce NA, Ciph A {|Nonce NA, Key K|},
Ciph B {|Nonce NB, Key K|}|}"
- or3':: "agent => msg => agent => agent => nat => nat => key => event"
+abbreviation
+ or3':: "agent => msg => agent => agent => nat => nat => key => event" where
"or3' S Y A B NA NB K ==
Says S B {|Nonce NA, Y, Ciph B {|Nonce NB, Key K|}|}"
- or4 :: "agent => agent => nat => msg => event"
+abbreviation
+ or4 :: "agent => agent => nat => msg => event" where
"or4 A B NA X == Says B A {|Nonce NA, X, nil|}"
- or4' :: "agent => agent => nat => key => event"
+abbreviation
+ or4' :: "agent => agent => nat => key => event" where
"or4' B' A NA K == Says B' A {|Nonce NA, Ciph A {|Nonce NA, Key K|}, nil|}"
subsection{*definition of the protocol*}
--- a/src/HOL/Auth/Guard/Guard_Shared.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Guard/Guard_Shared.thy Fri Nov 17 02:20:03 2006 +0100
@@ -20,7 +20,7 @@
subsubsection{*a little abbreviation*}
abbreviation
- Ciph :: "agent => msg => msg"
+ Ciph :: "agent => msg => msg" where
"Ciph A X == Crypt (shrK A) X"
subsubsection{*agent associated to a key*}
--- a/src/HOL/Auth/Guard/Guard_Yahalom.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Guard/Guard_Yahalom.thy Fri Nov 17 02:20:03 2006 +0100
@@ -16,31 +16,38 @@
subsection{*messages used in the protocol*}
abbreviation (input)
- ya1 :: "agent => agent => nat => event"
+ ya1 :: "agent => agent => nat => event" where
"ya1 A B NA == Says A B {|Agent A, Nonce NA|}"
- ya1' :: "agent => agent => agent => nat => event"
+abbreviation (input)
+ ya1' :: "agent => agent => agent => nat => event" where
"ya1' A' A B NA == Says A' B {|Agent A, Nonce NA|}"
- ya2 :: "agent => agent => nat => nat => event"
+abbreviation (input)
+ ya2 :: "agent => agent => nat => nat => event" where
"ya2 A B NA NB == Says B Server {|Agent B, Ciph B {|Agent A, Nonce NA, Nonce NB|}|}"
- ya2' :: "agent => agent => agent => nat => nat => event"
+abbreviation (input)
+ ya2' :: "agent => agent => agent => nat => nat => event" where
"ya2' B' A B NA NB == Says B' Server {|Agent B, Ciph B {|Agent A, Nonce NA, Nonce NB|}|}"
- ya3 :: "agent => agent => nat => nat => key => event"
+abbreviation (input)
+ ya3 :: "agent => agent => nat => nat => key => event" where
"ya3 A B NA NB K ==
Says Server A {|Ciph A {|Agent B, Key K, Nonce NA, Nonce NB|},
Ciph B {|Agent A, Key K|}|}"
- ya3':: "agent => msg => agent => agent => nat => nat => key => event"
+abbreviation (input)
+ ya3':: "agent => msg => agent => agent => nat => nat => key => event" where
"ya3' S Y A B NA NB K ==
Says S A {|Ciph A {|Agent B, Key K, Nonce NA, Nonce NB|}, Y|}"
- ya4 :: "agent => agent => nat => nat => msg => event"
+abbreviation (input)
+ ya4 :: "agent => agent => nat => nat => msg => event" where
"ya4 A B K NB Y == Says A B {|Y, Crypt K (Nonce NB)|}"
- ya4' :: "agent => agent => nat => nat => msg => event"
+abbreviation (input)
+ ya4' :: "agent => agent => nat => nat => msg => event" where
"ya4' A' B K NB Y == Says A' B {|Y, Crypt K (Nonce NB)|}"
--- a/src/HOL/Auth/Guard/List_Msg.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Guard/List_Msg.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,7 +18,7 @@
subsubsection{*nil is represented by any message which is not a pair*}
abbreviation (input)
- cons :: "msg => msg => msg"
+ cons :: "msg => msg => msg" where
"cons x l == {|x,l|}"
subsubsection{*induction principle*}
@@ -134,7 +134,7 @@
subsubsection{*set of well-formed agent-list messages*}
abbreviation
- nil :: msg
+ nil :: msg where
"nil == Number 0"
consts agl :: "msg set"
--- a/src/HOL/Auth/Guard/Proto.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Guard/Proto.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,7 +18,7 @@
types rule = "event set * event"
abbreviation
- msg' :: "rule => msg"
+ msg' :: "rule => msg" where
"msg' R == msg (snd R)"
types proto = "rule set"
@@ -77,16 +77,19 @@
"ap s (Notes A X) = Notes (agent s A) (apm s X)"
abbreviation
- ap' :: "subs => rule => event"
+ ap' :: "subs => rule => event" where
"ap' s R == ap s (snd R)"
- apm' :: "subs => rule => msg"
+abbreviation
+ apm' :: "subs => rule => msg" where
"apm' s R == apm s (msg' R)"
- priK' :: "subs => agent => key"
+abbreviation
+ priK' :: "subs => agent => key" where
"priK' s A == priK (agent s A)"
- pubK' :: "subs => agent => key"
+abbreviation
+ pubK' :: "subs => agent => key" where
"pubK' s A == pubK (agent s A)"
subsection{*nonces generated by a rule*}
@@ -380,14 +383,16 @@
ns :: proto
abbreviation
- ns1 :: rule
+ ns1 :: rule where
"ns1 == ({}, Says a b (Crypt (pubK b) {|Nonce Na, Agent a|}))"
- ns2 :: rule
+abbreviation
+ ns2 :: rule where
"ns2 == ({Says a' b (Crypt (pubK b) {|Nonce Na, Agent a|})},
Says b a (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|}))"
- ns3 :: rule
+abbreviation
+ ns3 :: rule where
"ns3 == ({Says a b (Crypt (pubK b) {|Nonce Na, Agent a|}),
Says b' a (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|})},
Says a b (Crypt (pubK b) (Nonce Nb)))"
@@ -398,10 +403,11 @@
[iff]: "ns3:ns"
abbreviation (input)
- ns3a :: event
+ ns3a :: event where
"ns3a == Says a b (Crypt (pubK b) {|Nonce Na, Agent a|})"
- ns3b :: event
+abbreviation (input)
+ ns3b :: event where
"ns3b == Says b' a (Crypt (pubK a) {|Nonce Na, Nonce Nb, Agent b|})"
constdefs keys :: "keyfun"
--- a/src/HOL/Auth/KerberosIV.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/KerberosIV.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,11 +11,10 @@
text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
abbreviation
- Kas :: agent
- "Kas == Server"
+ Kas :: agent where "Kas == Server"
- Tgs :: agent
- "Tgs == Friend 0"
+abbreviation
+ Tgs :: agent where "Tgs == Friend 0"
axioms
@@ -79,19 +78,23 @@
abbreviation
(*The current time is the length of the trace*)
- CT :: "event list=>nat"
+ CT :: "event list=>nat" where
"CT == length"
- expiredAK :: "[nat, event list] => bool"
+abbreviation
+ expiredAK :: "[nat, event list] => bool" where
"expiredAK Ta evs == authKlife + Ta < CT evs"
- expiredSK :: "[nat, event list] => bool"
+abbreviation
+ expiredSK :: "[nat, event list] => bool" where
"expiredSK Ts evs == servKlife + Ts < CT evs"
- expiredA :: "[nat, event list] => bool"
+abbreviation
+ expiredA :: "[nat, event list] => bool" where
"expiredA T evs == authlife + T < CT evs"
- valid :: "[nat, nat] => bool" ("valid _ wrt _")
+abbreviation
+ valid :: "[nat, nat] => bool" ("valid _ wrt _") where
"valid T1 wrt T2 == T1 <= replylife + T2"
(*---------------------------------------------------------------------*)
--- a/src/HOL/Auth/KerberosIV_Gets.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/KerberosIV_Gets.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,11 +11,10 @@
text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
abbreviation
- Kas :: agent
- "Kas == Server"
+ Kas :: agent where "Kas == Server"
- Tgs :: agent
- "Tgs == Friend 0"
+abbreviation
+ Tgs :: agent where "Tgs == Friend 0"
axioms
@@ -67,19 +66,23 @@
abbreviation
(*The current time is just the length of the trace!*)
- CT :: "event list=>nat"
+ CT :: "event list=>nat" where
"CT == length"
- expiredAK :: "[nat, event list] => bool"
+abbreviation
+ expiredAK :: "[nat, event list] => bool" where
"expiredAK Ta evs == authKlife + Ta < CT evs"
- expiredSK :: "[nat, event list] => bool"
+abbreviation
+ expiredSK :: "[nat, event list] => bool" where
"expiredSK Ts evs == servKlife + Ts < CT evs"
- expiredA :: "[nat, event list] => bool"
+abbreviation
+ expiredA :: "[nat, event list] => bool" where
"expiredA T evs == authlife + T < CT evs"
- valid :: "[nat, nat] => bool" ("valid _ wrt _")
+abbreviation
+ valid :: "[nat, nat] => bool" ("valid _ wrt _") where
"valid T1 wrt T2 == T1 <= replylife + T2"
(*---------------------------------------------------------------------*)
--- a/src/HOL/Auth/KerberosV.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/KerberosV.thy Fri Nov 17 02:20:03 2006 +0100
@@ -9,10 +9,11 @@
text{*The "u" prefix indicates theorems referring to an updated version of the protocol. The "r" suffix indicates theorems where the confidentiality assumptions are relaxed by the corresponding arguments.*}
abbreviation
- Kas :: agent
+ Kas :: agent where
"Kas == Server"
- Tgs :: agent
+abbreviation
+ Tgs :: agent where
"Tgs == Friend 0"
@@ -68,19 +69,23 @@
abbreviation
(*The current time is just the length of the trace!*)
- CT :: "event list=>nat"
+ CT :: "event list=>nat" where
"CT == length"
- expiredAK :: "[nat, event list] => bool"
+abbreviation
+ expiredAK :: "[nat, event list] => bool" where
"expiredAK T evs == authKlife + T < CT evs"
- expiredSK :: "[nat, event list] => bool"
+abbreviation
+ expiredSK :: "[nat, event list] => bool" where
"expiredSK T evs == servKlife + T < CT evs"
- expiredA :: "[nat, event list] => bool"
+abbreviation
+ expiredA :: "[nat, event list] => bool" where
"expiredA T evs == authlife + T < CT evs"
- valid :: "[nat, nat] => bool" ("valid _ wrt _")
+abbreviation
+ valid :: "[nat, nat] => bool" ("valid _ wrt _") where
"valid T1 wrt T2 == T1 <= replylife + T2"
(*---------------------------------------------------------------------*)
--- a/src/HOL/Auth/Kerberos_BAN.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Kerberos_BAN.thy Fri Nov 17 02:20:03 2006 +0100
@@ -39,13 +39,15 @@
by blast
abbreviation
- CT :: "event list=>nat"
+ CT :: "event list=>nat" where
"CT == length "
- expiredK :: "[nat, event list] => bool"
+abbreviation
+ expiredK :: "[nat, event list] => bool" where
"expiredK T evs == sesKlife + T < CT evs"
- expiredA :: "[nat, event list] => bool"
+abbreviation
+ expiredA :: "[nat, event list] => bool" where
"expiredA T evs == authlife + T < CT evs"
--- a/src/HOL/Auth/Kerberos_BAN_Gets.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Kerberos_BAN_Gets.thy Fri Nov 17 02:20:03 2006 +0100
@@ -40,13 +40,15 @@
abbreviation
- CT :: "event list=>nat"
+ CT :: "event list=>nat" where
"CT == length"
- expiredK :: "[nat, event list] => bool"
+abbreviation
+ expiredK :: "[nat, event list] => bool" where
"expiredK T evs == sesKlife + T < CT evs"
- expiredA :: "[nat, event list] => bool"
+abbreviation
+ expiredA :: "[nat, event list] => bool" where
"expiredA T evs == authlife + T < CT evs"
--- a/src/HOL/Auth/Public.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Public.thy Fri Nov 17 02:20:03 2006 +0100
@@ -21,19 +21,24 @@
publicKey :: "[keymode,agent] => key"
abbreviation
- pubEK :: "agent => key"
+ pubEK :: "agent => key" where
"pubEK == publicKey Encryption"
- pubSK :: "agent => key"
+abbreviation
+ pubSK :: "agent => key" where
"pubSK == publicKey Signature"
- privateKey :: "[keymode, agent] => key"
+abbreviation
+ privateKey :: "[keymode, agent] => key" where
"privateKey b A == invKey (publicKey b A)"
+abbreviation
(*BEWARE!! priEK, priSK DON'T WORK with inj, range, image, etc.*)
- priEK :: "agent => key"
+ priEK :: "agent => key" where
"priEK A == privateKey Encryption A"
- priSK :: "agent => key"
+
+abbreviation
+ priSK :: "agent => key" where
"priSK A == privateKey Signature A"
@@ -41,10 +46,11 @@
simple situation where the signature and encryption keys are the same.*}
abbreviation
- pubK :: "agent => key"
+ pubK :: "agent => key" where
"pubK A == pubEK A"
- priK :: "agent => key"
+abbreviation
+ priK :: "agent => key" where
"priK A == invKey (pubEK A)"
--- a/src/HOL/Auth/Recur.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Recur.thy Fri Nov 17 02:20:03 2006 +0100
@@ -10,7 +10,7 @@
text{*End marker for message bundles*}
abbreviation
- END :: "msg"
+ END :: "msg" where
"END == Number 0"
(*Two session keys are distributed to each agent except for the initiator,
--- a/src/HOL/Auth/Smartcard/EventSC.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/Smartcard/EventSC.thy Fri Nov 17 02:20:03 2006 +0100
@@ -25,7 +25,7 @@
secureM :: "bool"(*assumption of secure means between agents and their cards*)
abbreviation
- insecureM :: bool (*certain protocols make no assumption of secure means*)
+ insecureM :: bool where (*certain protocols make no assumption of secure means*)
"insecureM == \<not>secureM"
--- a/src/HOL/Auth/TLS.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/TLS.thy Fri Nov 17 02:20:03 2006 +0100
@@ -64,10 +64,11 @@
sessionK :: "(nat*nat*nat) * role => key"
abbreviation
- clientK :: "nat*nat*nat => key"
+ clientK :: "nat*nat*nat => key" where
"clientK X == sessionK(X, ClientRole)"
- serverK :: "nat*nat*nat => key"
+abbreviation
+ serverK :: "nat*nat*nat => key" where
"serverK X == sessionK(X, ServerRole)"
--- a/src/HOL/Auth/ZhouGollmann.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Auth/ZhouGollmann.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,17 +13,12 @@
theory ZhouGollmann imports Public begin
abbreviation
- TTP :: agent
- "TTP == Server"
+ TTP :: agent where "TTP == Server"
- f_sub :: nat
- "f_sub == 5"
- f_nro :: nat
- "f_nro == 2"
- f_nrr :: nat
- "f_nrr == 3"
- f_con :: nat
- "f_con == 4"
+abbreviation f_sub :: nat where "f_sub == 5"
+abbreviation f_nro :: nat where "f_nro == 2"
+abbreviation f_nrr :: nat where "f_nrr == 3"
+abbreviation f_con :: nat where "f_con == 4"
constdefs
--- a/src/HOL/Bali/Example.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Bali/Example.thy Fri Nov 17 02:20:03 2006 +0100
@@ -87,31 +87,39 @@
surj_label_:" \<exists>m. n = label_ m"
abbreviation
- HasFoo :: qtname
+ HasFoo :: qtname where
"HasFoo == \<lparr>pid=java_lang,tid=TName (tnam_ HasFoo_)\<rparr>"
- Base :: qtname
+abbreviation
+ Base :: qtname where
"Base == \<lparr>pid=java_lang,tid=TName (tnam_ Base_)\<rparr>"
- Ext :: qtname
+abbreviation
+ Ext :: qtname where
"Ext == \<lparr>pid=java_lang,tid=TName (tnam_ Ext_)\<rparr>"
- Main :: qtname
+abbreviation
+ Main :: qtname where
"Main == \<lparr>pid=java_lang,tid=TName (tnam_ Main_)\<rparr>"
- arr :: vname
+abbreviation
+ arr :: vname where
"arr == (vnam_ arr_)"
- vee :: vname
+abbreviation
+ vee :: vname where
"vee == (vnam_ vee_)"
- z :: vname
+abbreviation
+ z :: vname where
"z == (vnam_ z_)"
- e :: vname
+abbreviation
+ e :: vname where
"e == (vnam_ e_)"
- lab1:: label
+abbreviation
+ lab1:: label where
"lab1 == label_ lab1_"
@@ -261,7 +269,7 @@
section "program"
abbreviation
- tprg :: prog
+ tprg :: prog where
"tprg == \<lparr>ifaces=Ifaces,classes=Classes\<rparr>"
constdefs
@@ -1195,11 +1203,10 @@
b :: loc
c :: loc
-abbreviation
- "one == Suc 0"
- "two == Suc one"
- "tree == Suc two"
- "four == Suc tree"
+abbreviation "one == Suc 0"
+abbreviation "two == Suc one"
+abbreviation "tree == Suc two"
+abbreviation "four == Suc tree"
syntax
obj_a :: obj
--- a/src/HOL/Code_Generator.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Code_Generator.thy Fri Nov 17 02:20:03 2006 +0100
@@ -175,7 +175,7 @@
text {* lazy @{const If} *}
definition
- if_delayed :: "bool \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> 'a"
+ if_delayed :: "bool \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> 'a" where
"if_delayed b f g = (if b then f True else g False)"
lemma [code func]:
--- a/src/HOL/Complex/CLim.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Complex/CLim.thy Fri Nov 17 02:20:03 2006 +0100
@@ -34,43 +34,50 @@
done
abbreviation
-
CLIM :: "[complex=>complex,complex,complex] => bool"
- ("((_)/ -- (_)/ --C> (_))" [60, 0, 60] 60)
+ ("((_)/ -- (_)/ --C> (_))" [60, 0, 60] 60) where
"CLIM == LIM"
+abbreviation
NSCLIM :: "[complex=>complex,complex,complex] => bool"
- ("((_)/ -- (_)/ --NSC> (_))" [60, 0, 60] 60)
+ ("((_)/ -- (_)/ --NSC> (_))" [60, 0, 60] 60) where
"NSCLIM == NSLIM"
+abbreviation
(* f: C --> R *)
CRLIM :: "[complex=>real,complex,real] => bool"
- ("((_)/ -- (_)/ --CR> (_))" [60, 0, 60] 60)
+ ("((_)/ -- (_)/ --CR> (_))" [60, 0, 60] 60) where
"CRLIM == LIM"
+abbreviation
NSCRLIM :: "[complex=>real,complex,real] => bool"
- ("((_)/ -- (_)/ --NSCR> (_))" [60, 0, 60] 60)
+ ("((_)/ -- (_)/ --NSCR> (_))" [60, 0, 60] 60) where
"NSCRLIM == NSLIM"
-
- isContc :: "[complex=>complex,complex] => bool"
+abbreviation
+ isContc :: "[complex=>complex,complex] => bool" where
"isContc == isCont"
+abbreviation
(* NS definition dispenses with limit notions *)
- isNSContc :: "[complex=>complex,complex] => bool"
+ isNSContc :: "[complex=>complex,complex] => bool" where
"isNSContc == isNSCont"
- isContCR :: "[complex=>real,complex] => bool"
+abbreviation
+ isContCR :: "[complex=>real,complex] => bool" where
"isContCR == isCont"
+abbreviation
(* NS definition dispenses with limit notions *)
- isNSContCR :: "[complex=>real,complex] => bool"
+ isNSContCR :: "[complex=>real,complex] => bool" where
"isNSContCR == isNSCont"
- isUContc :: "(complex=>complex) => bool"
+abbreviation
+ isUContc :: "(complex=>complex) => bool" where
"isUContc == isUCont"
- isNSUContc :: "(complex=>complex) => bool"
+abbreviation
+ isNSUContc :: "(complex=>complex) => bool" where
"isNSUContc == isNSUCont"
@@ -129,25 +136,27 @@
by (rule isNSUCont_def)
+ (* differentiation: D is derivative of function f at x *)
definition
-
- (* differentiation: D is derivative of function f at x *)
cderiv:: "[complex=>complex,complex,complex] => bool"
- ("(CDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
+ ("(CDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) where
"CDERIV f x :> D = ((%h. (f(x + h) - f(x))/h) -- 0 --C> D)"
+definition
nscderiv :: "[complex=>complex,complex,complex] => bool"
- ("(NSCDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60)
+ ("(NSCDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) where
"NSCDERIV f x :> D = (\<forall>h \<in> Infinitesimal - {0}.
(( *f* f)(hcomplex_of_complex x + h)
- hcomplex_of_complex (f x))/h @= hcomplex_of_complex D)"
+definition
cdifferentiable :: "[complex=>complex,complex] => bool"
- (infixl "cdifferentiable" 60)
+ (infixl "cdifferentiable" 60) where
"f cdifferentiable x = (\<exists>D. CDERIV f x :> D)"
+definition
NSCdifferentiable :: "[complex=>complex,complex] => bool"
- (infixl "NSCdifferentiable" 60)
+ (infixl "NSCdifferentiable" 60) where
"f NSCdifferentiable x = (\<exists>D. NSCDERIV f x :> D)"
--- a/src/HOL/Complex/Complex.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Complex/Complex.thy Fri Nov 17 02:20:03 2006 +0100
@@ -235,7 +235,7 @@
subsection{*Embedding Properties for @{term complex_of_real} Map*}
abbreviation
- complex_of_real :: "real => complex"
+ complex_of_real :: "real => complex" where
"complex_of_real == of_real"
lemma complex_of_real_def: "complex_of_real r = Complex r 0"
@@ -321,7 +321,7 @@
subsection{*Conjugation is an Automorphism*}
definition
- cnj :: "complex => complex"
+ cnj :: "complex => complex" where
"cnj z = Complex (Re z) (-Im z)"
lemma complex_cnj: "cnj (Complex x y) = Complex x (-y)"
@@ -385,7 +385,7 @@
complex_norm_def: "norm z == sqrt(Re(z) ^ 2 + Im(z) ^ 2)"
abbreviation
- cmod :: "complex => real"
+ cmod :: "complex => real" where
"cmod == norm"
lemmas cmod_def = complex_norm_def
@@ -575,10 +575,11 @@
definition
(*------------ Argand -------------*)
- sgn :: "complex => complex"
+ sgn :: "complex => complex" where
"sgn z = z / complex_of_real(cmod z)"
- arg :: "complex => real"
+definition
+ arg :: "complex => real" where
"arg z = (SOME a. Re(sgn z) = cos a & Im(sgn z) = sin a & -pi < a & a \<le> pi)"
lemma sgn_zero [simp]: "sgn 0 = 0"
@@ -671,15 +672,17 @@
definition
(* abbreviation for (cos a + i sin a) *)
- cis :: "real => complex"
+ cis :: "real => complex" where
"cis a = Complex (cos a) (sin a)"
+definition
(* abbreviation for r*(cos a + i sin a) *)
- rcis :: "[real, real] => complex"
+ rcis :: "[real, real] => complex" where
"rcis r a = complex_of_real r * cis a"
+definition
(* e ^ (x + iy) *)
- expi :: "complex => complex"
+ expi :: "complex => complex" where
"expi z = complex_of_real(exp (Re z)) * cis (Im z)"
lemma complex_split_polar:
--- a/src/HOL/Complex/NSCA.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Complex/NSCA.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,10 +11,11 @@
definition
(* standard complex numbers reagarded as an embedded subset of NS complex *)
- SComplex :: "hcomplex set"
+ SComplex :: "hcomplex set" where
"SComplex = {x. \<exists>r. x = hcomplex_of_complex r}"
- stc :: "hcomplex => hcomplex"
+definition
+ stc :: "hcomplex => hcomplex" where
--{* standard part map*}
"stc x = (SOME r. x \<in> HFinite & r:SComplex & r @= x)"
--- a/src/HOL/Complex/NSComplex.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Complex/NSComplex.thy Fri Nov 17 02:20:03 2006 +0100
@@ -14,63 +14,74 @@
types hcomplex = "complex star"
abbreviation
- hcomplex_of_complex :: "complex => complex star"
+ hcomplex_of_complex :: "complex => complex star" where
"hcomplex_of_complex == star_of"
- hcmod :: "complex star => real star"
+abbreviation
+ hcmod :: "complex star => real star" where
"hcmod == hnorm"
-definition
(*--- real and Imaginary parts ---*)
- hRe :: "hcomplex => hypreal"
+definition
+ hRe :: "hcomplex => hypreal" where
"hRe = *f* Re"
- hIm :: "hcomplex => hypreal"
+definition
+ hIm :: "hcomplex => hypreal" where
"hIm = *f* Im"
(*------ imaginary unit ----------*)
- iii :: hcomplex
+definition
+ iii :: hcomplex where
"iii = star_of ii"
(*------- complex conjugate ------*)
- hcnj :: "hcomplex => hcomplex"
+definition
+ hcnj :: "hcomplex => hcomplex" where
"hcnj = *f* cnj"
(*------------ Argand -------------*)
- hsgn :: "hcomplex => hcomplex"
+definition
+ hsgn :: "hcomplex => hcomplex" where
"hsgn = *f* sgn"
- harg :: "hcomplex => hypreal"
+definition
+ harg :: "hcomplex => hypreal" where
"harg = *f* arg"
+definition
(* abbreviation for (cos a + i sin a) *)
- hcis :: "hypreal => hcomplex"
+ hcis :: "hypreal => hcomplex" where
"hcis = *f* cis"
(*----- injection from hyperreals -----*)
- hcomplex_of_hypreal :: "hypreal => hcomplex"
+definition
+ hcomplex_of_hypreal :: "hypreal => hcomplex" where
"hcomplex_of_hypreal = *f* complex_of_real"
+definition
(* abbreviation for r*(cos a + i sin a) *)
- hrcis :: "[hypreal, hypreal] => hcomplex"
+ hrcis :: "[hypreal, hypreal] => hcomplex" where
"hrcis = *f2* rcis"
(*------------ e ^ (x + iy) ------------*)
-
- hexpi :: "hcomplex => hcomplex"
+definition
+ hexpi :: "hcomplex => hcomplex" where
"hexpi = *f* expi"
- HComplex :: "[hypreal,hypreal] => hcomplex"
+definition
+ HComplex :: "[hypreal,hypreal] => hcomplex" where
"HComplex = *f2* Complex"
- hcpow :: "[hcomplex,hypnat] => hcomplex" (infixr "hcpow" 80)
+definition
+ hcpow :: "[hcomplex,hypnat] => hcomplex" (infixr "hcpow" 80) where
"(z::hcomplex) hcpow (n::hypnat) = ( *f2* op ^) z n"
lemmas hcomplex_defs [transfer_unfold] =
--- a/src/HOL/Complex/ex/NSPrimes.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Complex/ex/NSPrimes.thy Fri Nov 17 02:20:03 2006 +0100
@@ -14,13 +14,15 @@
primes by considering a property of nonstandard sets.*}
definition
- hdvd :: "[hypnat, hypnat] => bool" (infixl "hdvd" 50)
+ hdvd :: "[hypnat, hypnat] => bool" (infixl "hdvd" 50) where
[transfer_unfold]: "(M::hypnat) hdvd N = ( *p2* (op dvd)) M N"
- starprime :: "hypnat set"
+definition
+ starprime :: "hypnat set" where
[transfer_unfold]: "starprime = ( *s* {p. prime p})"
- choicefun :: "'a set => 'a"
+definition
+ choicefun :: "'a set => 'a" where
"choicefun E = (@x. \<exists>X \<in> Pow(E) -{{}}. x : X)"
consts injf_max :: "nat => ('a::{order} set) => 'a"
--- a/src/HOL/Complex/ex/Sqrt.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Complex/ex/Sqrt.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,7 +18,7 @@
*}
definition
- rationals ("\<rat>")
+ rationals ("\<rat>") where
"\<rat> = {x. \<exists>m n. n \<noteq> 0 \<and> \<bar>x\<bar> = real (m::nat) / real (n::nat)}"
theorem rationals_rep [elim?]:
--- a/src/HOL/Complex/ex/Sqrt_Script.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Complex/ex/Sqrt_Script.thy Fri Nov 17 02:20:03 2006 +0100
@@ -53,7 +53,7 @@
subsection {* The set of rational numbers *}
definition
- rationals :: "real set" ("\<rat>")
+ rationals :: "real set" ("\<rat>") where
"\<rat> = {x. \<exists>m n. n \<noteq> 0 \<and> \<bar>x\<bar> = real (m::nat) / real (n::nat)}"
--- a/src/HOL/Datatype.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Datatype.thy Fri Nov 17 02:20:03 2006 +0100
@@ -726,7 +726,7 @@
subsubsection {* Code generator setup *}
definition
- is_none :: "'a option \<Rightarrow> bool"
+ is_none :: "'a option \<Rightarrow> bool" where
is_none_none [normal post, symmetric, code inline]: "is_none x \<longleftrightarrow> x = None"
lemma is_none_code [code]:
--- a/src/HOL/Equiv_Relations.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Equiv_Relations.thy Fri Nov 17 02:20:03 2006 +0100
@@ -164,7 +164,8 @@
assumes congruent: "(y,z) \<in> r ==> f y = f z"
abbreviation
- RESPECTS :: "('a => 'b) => ('a * 'a) set => bool" (infixr "respects" 80)
+ RESPECTS :: "('a => 'b) => ('a * 'a) set => bool"
+ (infixr "respects" 80) where
"f respects r == congruent r f"
@@ -222,7 +223,8 @@
text{*Abbreviation for the common case where the relations are identical*}
abbreviation
- RESPECTS2:: "['a => 'a => 'b, ('a * 'a)set] => bool" (infixr "respects2 " 80)
+ RESPECTS2:: "['a => 'a => 'b, ('a * 'a) set] => bool"
+ (infixr "respects2 " 80) where
"f respects2 r == congruent2 r r f"
--- a/src/HOL/Extraction/Pigeonhole.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Extraction/Pigeonhole.thy Fri Nov 17 02:20:03 2006 +0100
@@ -302,9 +302,10 @@
ML "PH.pigeonhole 8 (PH.sel [0,1,2,3,4,5,6,3,7,8])"
definition
- arbitrary_nat :: "nat \<times> nat"
+ arbitrary_nat :: "nat \<times> nat" where
[symmetric, code inline]: "arbitrary_nat = arbitrary"
- arbitrary_nat_subst :: "nat \<times> nat"
+definition
+ arbitrary_nat_subst :: "nat \<times> nat" where
"arbitrary_nat_subst = (0, 0)"
code_axioms
@@ -312,6 +313,7 @@
definition
"test n = pigeonhole n (\<lambda>m. m - 1)"
+definition
"test' n = pigeonhole_slow n (\<lambda>m. m - 1)"
code_gen test test' "op !" (SML *)
--- a/src/HOL/Finite_Set.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Finite_Set.thy Fri Nov 17 02:20:03 2006 +0100
@@ -803,7 +803,7 @@
"setsum f A == if finite A then fold (op +) f 0 A else 0"
abbreviation
- Setsum ("\<Sum>_" [1000] 999)
+ Setsum ("\<Sum>_" [1000] 999) where
"\<Sum>A == setsum (%x. x) A"
text{* Now: lot's of fancy syntax. First, @{term "setsum (%x. e) A"} is
@@ -1275,7 +1275,7 @@
"setprod f A == if finite A then fold (op *) f 1 A else 1"
abbreviation
- Setprod ("\<Prod>_" [1000] 999)
+ Setprod ("\<Prod>_" [1000] 999) where
"\<Prod>A == setprod (%x. x) A"
syntax
--- a/src/HOL/FixedPoint.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/FixedPoint.thy Fri Nov 17 02:20:03 2006 +0100
@@ -20,12 +20,13 @@
defs Sup_def: "Sup A == Meet {b. \<forall>a \<in> A. a <= b}"
definition
- SUP :: "('a \<Rightarrow> 'b::order) \<Rightarrow> 'b" (binder "SUP " 10)
-"SUP x. f x == Sup (f ` UNIV)"
+ SUP :: "('a \<Rightarrow> 'b::order) \<Rightarrow> 'b" (binder "SUP " 10) where
+ "SUP x. f x == Sup (f ` UNIV)"
+
(*
abbreviation
- bot :: "'a::order"
-"bot == Sup {}"
+ bot :: "'a::order" where
+ "bot == Sup {}"
*)
axclass comp_lat < order
Meet_lower: "x \<in> A \<Longrightarrow> Meet A <= x"
--- a/src/HOL/FunDef.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/FunDef.thy Fri Nov 17 02:20:03 2006 +0100
@@ -108,7 +108,7 @@
section {* Definitions with default value *}
definition
- THE_default :: "'a \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a"
+ THE_default :: "'a \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a" where
"THE_default d P = (if (\<exists>!x. P x) then (THE x. P x) else d)"
lemma THE_defaultI': "\<exists>!x. P x \<Longrightarrow> P (THE_default d P)"
@@ -178,9 +178,9 @@
inductive rpg
intros
"(Inr y, y) : rpg"
-definition
- "lproj x = (THE y. (x,y) : lpg)"
- "rproj x = (THE y. (x,y) : rpg)"
+
+definition "lproj x = (THE y. (x,y) : lpg)"
+definition "rproj x = (THE y. (x,y) : rpg)"
lemma lproj_inl:
"lproj (Inl x) = x"
--- a/src/HOL/HOL.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/HOL.thy Fri Nov 17 02:20:03 2006 +0100
@@ -58,27 +58,27 @@
"op =" (infix "=" 50)
abbreviation
- not_equal :: "['a, 'a] => bool" (infixl "~=" 50)
+ not_equal :: "['a, 'a] => bool" (infixl "~=" 50) where
"x ~= y == ~ (x = y)"
notation (output)
not_equal (infix "~=" 50)
notation (xsymbols)
- Not ("\<not> _" [40] 40)
- "op &" (infixr "\<and>" 35)
- "op |" (infixr "\<or>" 30)
- "op -->" (infixr "\<longrightarrow>" 25)
+ Not ("\<not> _" [40] 40) and
+ "op &" (infixr "\<and>" 35) and
+ "op |" (infixr "\<or>" 30) and
+ "op -->" (infixr "\<longrightarrow>" 25) and
not_equal (infix "\<noteq>" 50)
notation (HTML output)
- Not ("\<not> _" [40] 40)
- "op &" (infixr "\<and>" 35)
- "op |" (infixr "\<or>" 30)
+ Not ("\<not> _" [40] 40) and
+ "op &" (infixr "\<and>" 35) and
+ "op |" (infixr "\<or>" 30) and
not_equal (infix "\<noteq>" 50)
abbreviation (iff)
- iff :: "[bool, bool] => bool" (infixr "<->" 25)
+ iff :: "[bool, bool] => bool" (infixr "<->" 25) where
"A <-> B == A = B"
notation (xsymbols)
--- a/src/HOL/Hyperreal/Deriv.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/Deriv.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,26 +15,29 @@
text{*Standard and Nonstandard Definitions*}
definition
-
deriv :: "[real \<Rightarrow> 'a::real_normed_vector, real, 'a] \<Rightarrow> bool"
--{*Differentiation: D is derivative of function f at x*}
- ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
+ ("(DERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
"DERIV f x :> D = ((%h. (f(x + h) - f x) /# h) -- 0 --> D)"
+definition
nsderiv :: "[real=>real,real,real] => bool"
- ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
+ ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
"NSDERIV f x :> D = (\<forall>h \<in> Infinitesimal - {0}.
(( *f* f)(hypreal_of_real x + h)
- hypreal_of_real (f x))/h @= hypreal_of_real D)"
- differentiable :: "[real=>real,real] => bool" (infixl "differentiable" 60)
+definition
+ differentiable :: "[real=>real,real] => bool" (infixl "differentiable" 60) where
"f differentiable x = (\<exists>D. DERIV f x :> D)"
+definition
NSdifferentiable :: "[real=>real,real] => bool"
- (infixl "NSdifferentiable" 60)
+ (infixl "NSdifferentiable" 60) where
"f NSdifferentiable x = (\<exists>D. NSDERIV f x :> D)"
- increment :: "[real=>real,real,hypreal] => hypreal"
+definition
+ increment :: "[real=>real,real,hypreal] => hypreal" where
"increment f x h = (@inc. f NSdifferentiable x &
inc = ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x))"
--- a/src/HOL/Hyperreal/HLog.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/HLog.thy Fri Nov 17 02:20:03 2006 +0100
@@ -19,10 +19,11 @@
definition
- powhr :: "[hypreal,hypreal] => hypreal" (infixr "powhr" 80)
+ powhr :: "[hypreal,hypreal] => hypreal" (infixr "powhr" 80) where
"x powhr a = starfun2 (op powr) x a"
- hlog :: "[hypreal,hypreal] => hypreal"
+definition
+ hlog :: "[hypreal,hypreal] => hypreal" where
"hlog a x = starfun2 log a x"
declare powhr_def [transfer_unfold]
--- a/src/HOL/Hyperreal/HSeries.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/HSeries.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,17 +12,20 @@
begin
definition
- sumhr :: "(hypnat * hypnat * (nat=>real)) => hypreal"
+ sumhr :: "(hypnat * hypnat * (nat=>real)) => hypreal" where
"sumhr =
(%(M,N,f). starfun2 (%m n. setsum f {m..<n}) M N)"
- NSsums :: "[nat=>real,real] => bool" (infixr "NSsums" 80)
+definition
+ NSsums :: "[nat=>real,real] => bool" (infixr "NSsums" 80) where
"f NSsums s = (%n. setsum f {0..<n}) ----NS> s"
- NSsummable :: "(nat=>real) => bool"
+definition
+ NSsummable :: "(nat=>real) => bool" where
"NSsummable f = (\<exists>s. f NSsums s)"
- NSsuminf :: "(nat=>real) => real"
+definition
+ NSsuminf :: "(nat=>real) => real" where
"NSsuminf f = (THE s. f NSsums s)"
--- a/src/HOL/Hyperreal/HTranscendental.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/HTranscendental.thy Fri Nov 17 02:20:03 2006 +0100
@@ -31,15 +31,17 @@
definition
- exphr :: "real => hypreal"
+ exphr :: "real => hypreal" where
--{*define exponential function using standard part *}
"exphr x = st(sumhr (0, whn, %n. inverse(real (fact n)) * (x ^ n)))"
- sinhr :: "real => hypreal"
+definition
+ sinhr :: "real => hypreal" where
"sinhr x = st(sumhr (0, whn, %n. (if even(n) then 0 else
((-1) ^ ((n - 1) div 2))/(real (fact n))) * (x ^ n)))"
- coshr :: "real => hypreal"
+definition
+ coshr :: "real => hypreal" where
"coshr x = st(sumhr (0, whn, %n. (if even(n) then
((-1) ^ (n div 2))/(real (fact n)) else 0) * x ^ n))"
--- a/src/HOL/Hyperreal/HyperArith.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/HyperArith.thy Fri Nov 17 02:20:03 2006 +0100
@@ -34,7 +34,7 @@
subsection{*Embedding the Naturals into the Hyperreals*}
abbreviation
- hypreal_of_nat :: "nat => hypreal"
+ hypreal_of_nat :: "nat => hypreal" where
"hypreal_of_nat == of_nat"
lemma SNat_eq: "Nats = {n. \<exists>N. n = hypreal_of_nat N}"
--- a/src/HOL/Hyperreal/HyperDef.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/HyperDef.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,22 +15,23 @@
types hypreal = "real star"
abbreviation
- hypreal_of_real :: "real => real star"
+ hypreal_of_real :: "real => real star" where
"hypreal_of_real == star_of"
definition
- omega :: hypreal -- {*an infinite number @{text "= [<1,2,3,...>]"} *}
+ omega :: hypreal where -- {*an infinite number @{text "= [<1,2,3,...>]"} *}
"omega = star_n (%n. real (Suc n))"
- epsilon :: hypreal -- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *}
+definition
+ epsilon :: hypreal where -- {*an infinitesimal number @{text "= [<1,1/2,1/3,...>]"} *}
"epsilon = star_n (%n. inverse (real (Suc n)))"
notation (xsymbols)
- omega ("\<omega>")
+ omega ("\<omega>") and
epsilon ("\<epsilon>")
notation (HTML output)
- omega ("\<omega>")
+ omega ("\<omega>") and
epsilon ("\<epsilon>")
--- a/src/HOL/Hyperreal/HyperNat.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/HyperNat.thy Fri Nov 17 02:20:03 2006 +0100
@@ -14,7 +14,7 @@
types hypnat = "nat star"
abbreviation
- hypnat_of_nat :: "nat => nat star"
+ hypnat_of_nat :: "nat => nat star" where
"hypnat_of_nat == star_of"
subsection{*Properties Transferred from Naturals*}
@@ -161,7 +161,7 @@
definition
(* the set of infinite hypernatural numbers *)
- HNatInfinite :: "hypnat set"
+ HNatInfinite :: "hypnat set" where
"HNatInfinite = {n. n \<notin> Nats}"
lemma Nats_not_HNatInfinite_iff: "(x \<in> Nats) = (x \<notin> HNatInfinite)"
@@ -254,7 +254,7 @@
definition
(* omega is in fact an infinite hypernatural number = [<1,2,3,...>] *)
- whn :: hypnat
+ whn :: hypnat where
hypnat_omega_def: "whn = star_n (%n::nat. n)"
lemma hypnat_of_nat_neq_whn: "hypnat_of_nat n \<noteq> whn"
@@ -362,7 +362,7 @@
text{*Obtained using the nonstandard extension of the naturals*}
definition
- hypreal_of_hypnat :: "hypnat => hypreal"
+ hypreal_of_hypnat :: "hypnat => hypreal" where
"hypreal_of_hypnat = *f* real"
declare hypreal_of_hypnat_def [transfer_unfold]
--- a/src/HOL/Hyperreal/HyperPow.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/HyperPow.thy Fri Nov 17 02:20:03 2006 +0100
@@ -20,7 +20,7 @@
definition
(* hypernatural powers of hyperreals *)
- pow :: "[hypreal,hypnat] => hypreal" (infixr "pow" 80)
+ pow :: "[hypreal,hypnat] => hypreal" (infixr "pow" 80) where
hyperpow_def [transfer_unfold]:
"(R::hypreal) pow (N::hypnat) = ( *f2* op ^) R N"
--- a/src/HOL/Hyperreal/Integration.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/Integration.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,38 +13,43 @@
text{*We follow John Harrison in formalizing the Gauge integral.*}
definition
-
--{*Partitions and tagged partitions etc.*}
- partition :: "[(real*real),nat => real] => bool"
+ partition :: "[(real*real),nat => real] => bool" where
"partition = (%(a,b) D. D 0 = a &
(\<exists>N. (\<forall>n < N. D(n) < D(Suc n)) &
(\<forall>n \<ge> N. D(n) = b)))"
- psize :: "(nat => real) => nat"
+definition
+ psize :: "(nat => real) => nat" where
"psize D = (SOME N. (\<forall>n < N. D(n) < D(Suc n)) &
(\<forall>n \<ge> N. D(n) = D(N)))"
- tpart :: "[(real*real),((nat => real)*(nat =>real))] => bool"
+definition
+ tpart :: "[(real*real),((nat => real)*(nat =>real))] => bool" where
"tpart = (%(a,b) (D,p). partition(a,b) D &
(\<forall>n. D(n) \<le> p(n) & p(n) \<le> D(Suc n)))"
--{*Gauges and gauge-fine divisions*}
- gauge :: "[real => bool, real => real] => bool"
+definition
+ gauge :: "[real => bool, real => real] => bool" where
"gauge E g = (\<forall>x. E x --> 0 < g(x))"
- fine :: "[real => real, ((nat => real)*(nat => real))] => bool"
+definition
+ fine :: "[real => real, ((nat => real)*(nat => real))] => bool" where
"fine = (%g (D,p). \<forall>n. n < (psize D) --> D(Suc n) - D(n) < g(p n))"
--{*Riemann sum*}
- rsum :: "[((nat=>real)*(nat=>real)),real=>real] => real"
+definition
+ rsum :: "[((nat=>real)*(nat=>real)),real=>real] => real" where
"rsum = (%(D,p) f. \<Sum>n=0..<psize(D). f(p n) * (D(Suc n) - D(n)))"
--{*Gauge integrability (definite)*}
- Integral :: "[(real*real),real=>real,real] => bool"
+definition
+ Integral :: "[(real*real),real=>real,real] => bool" where
"Integral = (%(a,b) f k. \<forall>e > 0.
(\<exists>g. gauge(%x. a \<le> x & x \<le> b) g &
(\<forall>D p. tpart(a,b) (D,p) & fine(g)(D,p) -->
--- a/src/HOL/Hyperreal/Lim.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/Lim.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,28 +15,33 @@
definition
LIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool"
- ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60)
+ ("((_)/ -- (_)/ --> (_))" [60, 0, 60] 60) where
"f -- a --> L =
(\<forall>r > 0. \<exists>s > 0. \<forall>x. x \<noteq> a & norm (x - a) < s
--> norm (f x - L) < r)"
+definition
NSLIM :: "['a::real_normed_vector => 'b::real_normed_vector, 'a, 'b] => bool"
- ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60)
+ ("((_)/ -- (_)/ --NS> (_))" [60, 0, 60] 60) where
"f -- a --NS> L =
(\<forall>x. (x \<noteq> star_of a & x @= star_of a --> ( *f* f) x @= star_of L))"
- isCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool"
+definition
+ isCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where
"isCont f a = (f -- a --> (f a))"
- isNSCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool"
+definition
+ isNSCont :: "['a::real_normed_vector => 'b::real_normed_vector, 'a] => bool" where
--{*NS definition dispenses with limit notions*}
"isNSCont f a = (\<forall>y. y @= star_of a -->
( *f* f) y @= star_of (f a))"
- isUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool"
+definition
+ isUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where
"isUCont f = (\<forall>r>0. \<exists>s>0. \<forall>x y. norm (x - y) < s \<longrightarrow> norm (f x - f y) < r)"
- isNSUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool"
+definition
+ isNSUCont :: "['a::real_normed_vector => 'b::real_normed_vector] => bool" where
"isNSUCont f = (\<forall>x y. x @= y --> ( *f* f) x @= ( *f* f) y)"
--- a/src/HOL/Hyperreal/Log.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/Log.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,11 +11,12 @@
begin
definition
- powr :: "[real,real] => real" (infixr "powr" 80)
+ powr :: "[real,real] => real" (infixr "powr" 80) where
--{*exponentation with real exponent*}
"x powr a = exp(a * ln x)"
- log :: "[real,real] => real"
+definition
+ log :: "[real,real] => real" where
--{*logarithm of @{term x} to base @{term a}*}
"log a x = ln x / ln a"
--- a/src/HOL/Hyperreal/NSA.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/NSA.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,32 +12,37 @@
begin
definition
-
- hnorm :: "'a::norm star \<Rightarrow> real star"
+ hnorm :: "'a::norm star \<Rightarrow> real star" where
"hnorm = *f* norm"
- Infinitesimal :: "('a::real_normed_vector) star set"
+definition
+ Infinitesimal :: "('a::real_normed_vector) star set" where
"Infinitesimal = {x. \<forall>r \<in> Reals. 0 < r --> hnorm x < r}"
- HFinite :: "('a::real_normed_vector) star set"
+definition
+ HFinite :: "('a::real_normed_vector) star set" where
"HFinite = {x. \<exists>r \<in> Reals. hnorm x < r}"
- HInfinite :: "('a::real_normed_vector) star set"
+definition
+ HInfinite :: "('a::real_normed_vector) star set" where
"HInfinite = {x. \<forall>r \<in> Reals. r < hnorm x}"
- approx :: "['a::real_normed_vector star, 'a star] => bool"
- (infixl "@=" 50)
+definition
+ approx :: "['a::real_normed_vector star, 'a star] => bool" (infixl "@=" 50) where
--{*the `infinitely close' relation*}
"(x @= y) = ((x - y) \<in> Infinitesimal)"
- st :: "hypreal => hypreal"
+definition
+ st :: "hypreal => hypreal" where
--{*the standard part of a hyperreal*}
"st = (%x. @r. x \<in> HFinite & r \<in> Reals & r @= x)"
- monad :: "'a::real_normed_vector star => 'a star set"
+definition
+ monad :: "'a::real_normed_vector star => 'a star set" where
"monad x = {y. x @= y}"
- galaxy :: "'a::real_normed_vector star => 'a star set"
+definition
+ galaxy :: "'a::real_normed_vector star => 'a star set" where
"galaxy x = {y. (x + -y) \<in> HFinite}"
notation (xsymbols)
@@ -52,7 +57,7 @@
subsection {* Nonstandard Extension of the Norm Function *}
definition
- scaleHR :: "real star \<Rightarrow> 'a star \<Rightarrow> 'a::real_normed_vector star"
+ scaleHR :: "real star \<Rightarrow> 'a star \<Rightarrow> 'a::real_normed_vector star" where
"scaleHR = starfun2 scaleR"
declare hnorm_def [transfer_unfold]
--- a/src/HOL/Hyperreal/NatStar.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/NatStar.thy Fri Nov 17 02:20:03 2006 +0100
@@ -186,7 +186,7 @@
subsection{*Nonstandard Characterization of Induction*}
definition
- hSuc :: "hypnat => hypnat"
+ hSuc :: "hypnat => hypnat" where
"hSuc n = n + 1"
lemma starP: "(( *p* P) (star_n X)) = ({n. P (X n)} \<in> FreeUltrafilterNat)"
--- a/src/HOL/Hyperreal/NthRoot.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/NthRoot.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,11 +11,11 @@
begin
definition
-
- root :: "[nat, real] \<Rightarrow> real"
+ root :: "[nat, real] \<Rightarrow> real" where
"root n x = (THE u. (0 < x \<longrightarrow> 0 < u) \<and> (u ^ n = x))"
- sqrt :: "real \<Rightarrow> real"
+definition
+ sqrt :: "real \<Rightarrow> real" where
"sqrt x = root 2 x"
--- a/src/HOL/Hyperreal/Poly.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/Poly.thy Fri Nov 17 02:20:03 2006 +0100
@@ -82,25 +82,30 @@
text{*Other definitions*}
definition
- poly_minus :: "real list => real list" ("-- _" [80] 80)
+ poly_minus :: "real list => real list" ("-- _" [80] 80) where
"-- p = (- 1) %* p"
- pderiv :: "real list => real list"
+definition
+ pderiv :: "real list => real list" where
"pderiv p = (if p = [] then [] else pderiv_aux 1 (tl p))"
- divides :: "[real list,real list] => bool" (infixl "divides" 70)
+definition
+ divides :: "[real list,real list] => bool" (infixl "divides" 70) where
"p1 divides p2 = (\<exists>q. poly p2 = poly(p1 *** q))"
- order :: "real => real list => nat"
+definition
+ order :: "real => real list => nat" where
--{*order of a polynomial*}
"order a p = (SOME n. ([-a, 1] %^ n) divides p &
~ (([-a, 1] %^ (Suc n)) divides p))"
- degree :: "real list => nat"
+definition
+ degree :: "real list => nat" where
--{*degree of a polynomial*}
"degree p = length (pnormalize p)"
- rsquarefree :: "real list => bool"
+definition
+ rsquarefree :: "real list => bool" where
--{*squarefree polynomials --- NB with respect to real roots only.*}
"rsquarefree p = (poly p \<noteq> poly [] &
(\<forall>a. (order a p = 0) | (order a p = 1)))"
@@ -108,7 +113,7 @@
lemma padd_Nil2: "p +++ [] = p"
-by (induct "p", auto)
+by (induct p) auto
declare padd_Nil2 [simp]
lemma padd_Cons_Cons: "(h1 # p1) +++ (h2 # p2) = (h1 + h2) # (p1 +++ p2)"
--- a/src/HOL/Hyperreal/SEQ.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/SEQ.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,54 +13,64 @@
begin
definition
-
LIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
- ("((_)/ ----> (_))" [60, 60] 60)
+ ("((_)/ ----> (_))" [60, 60] 60) where
--{*Standard definition of convergence of sequence*}
"X ----> L = (\<forall>r. 0 < r --> (\<exists>no. \<forall>n. no \<le> n --> norm (X n - L) < r))"
+definition
NSLIMSEQ :: "[nat => 'a::real_normed_vector, 'a] => bool"
- ("((_)/ ----NS> (_))" [60, 60] 60)
+ ("((_)/ ----NS> (_))" [60, 60] 60) where
--{*Nonstandard definition of convergence of sequence*}
"X ----NS> L = (\<forall>N \<in> HNatInfinite. ( *f* X) N \<approx> star_of L)"
- lim :: "(nat => 'a::real_normed_vector) => 'a"
+definition
+ lim :: "(nat => 'a::real_normed_vector) => 'a" where
--{*Standard definition of limit using choice operator*}
"lim X = (THE L. X ----> L)"
- nslim :: "(nat => 'a::real_normed_vector) => 'a"
+definition
+ nslim :: "(nat => 'a::real_normed_vector) => 'a" where
--{*Nonstandard definition of limit using choice operator*}
"nslim X = (THE L. X ----NS> L)"
- convergent :: "(nat => 'a::real_normed_vector) => bool"
+definition
+ convergent :: "(nat => 'a::real_normed_vector) => bool" where
--{*Standard definition of convergence*}
"convergent X = (\<exists>L. X ----> L)"
- NSconvergent :: "(nat => 'a::real_normed_vector) => bool"
+definition
+ NSconvergent :: "(nat => 'a::real_normed_vector) => bool" where
--{*Nonstandard definition of convergence*}
"NSconvergent X = (\<exists>L. X ----NS> L)"
- Bseq :: "(nat => 'a::real_normed_vector) => bool"
+definition
+ Bseq :: "(nat => 'a::real_normed_vector) => bool" where
--{*Standard definition for bounded sequence*}
"Bseq X = (\<exists>K>0.\<forall>n. norm (X n) \<le> K)"
- NSBseq :: "(nat => 'a::real_normed_vector) => bool"
+definition
+ NSBseq :: "(nat => 'a::real_normed_vector) => bool" where
--{*Nonstandard definition for bounded sequence*}
"NSBseq X = (\<forall>N \<in> HNatInfinite. ( *f* X) N : HFinite)"
- monoseq :: "(nat=>real)=>bool"
+definition
+ monoseq :: "(nat=>real)=>bool" where
--{*Definition for monotonicity*}
"monoseq X = ((\<forall>m. \<forall>n\<ge>m. X m \<le> X n) | (\<forall>m. \<forall>n\<ge>m. X n \<le> X m))"
- subseq :: "(nat => nat) => bool"
+definition
+ subseq :: "(nat => nat) => bool" where
--{*Definition of subsequence*}
"subseq f = (\<forall>m. \<forall>n>m. (f m) < (f n))"
- Cauchy :: "(nat => 'a::real_normed_vector) => bool"
+definition
+ Cauchy :: "(nat => 'a::real_normed_vector) => bool" where
--{*Standard definition of the Cauchy condition*}
"Cauchy X = (\<forall>e>0. \<exists>M. \<forall>m \<ge> M. \<forall>n \<ge> M. norm (X m - X n) < e)"
- NSCauchy :: "(nat => 'a::real_normed_vector) => bool"
+definition
+ NSCauchy :: "(nat => 'a::real_normed_vector) => bool" where
--{*Nonstandard definition*}
"NSCauchy X = (\<forall>M \<in> HNatInfinite. \<forall>N \<in> HNatInfinite. ( *f* X) M \<approx> ( *f* X) N)"
--- a/src/HOL/Hyperreal/Series.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/Series.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,13 +15,15 @@
definition
sums :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool"
- (infixr "sums" 80)
+ (infixr "sums" 80) where
"f sums s = (%n. setsum f {0..<n}) ----> s"
- summable :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool"
+definition
+ summable :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> bool" where
"summable f = (\<exists>s. f sums s)"
- suminf :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a"
+definition
+ suminf :: "(nat \<Rightarrow> 'a::real_normed_vector) \<Rightarrow> 'a" where
"suminf f = (THE s. f sums s)"
syntax
--- a/src/HOL/Hyperreal/Star.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/Star.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,21 +12,26 @@
definition
(* internal sets *)
- starset_n :: "(nat => 'a set) => 'a star set" ("*sn* _" [80] 80)
+ starset_n :: "(nat => 'a set) => 'a star set" ("*sn* _" [80] 80) where
"*sn* As = Iset (star_n As)"
- InternalSets :: "'a star set set"
+definition
+ InternalSets :: "'a star set set" where
"InternalSets = {X. \<exists>As. X = *sn* As}"
+definition
(* nonstandard extension of function *)
- is_starext :: "['a star => 'a star, 'a => 'a] => bool"
+ is_starext :: "['a star => 'a star, 'a => 'a] => bool" where
"is_starext F f = (\<forall>x y. \<exists>X \<in> Rep_star(x). \<exists>Y \<in> Rep_star(y).
((y = (F x)) = ({n. Y n = f(X n)} : FreeUltrafilterNat)))"
+
+definition
(* internal functions *)
- starfun_n :: "(nat => ('a => 'b)) => 'a star => 'b star" ("*fn* _" [80] 80)
+ starfun_n :: "(nat => ('a => 'b)) => 'a star => 'b star" ("*fn* _" [80] 80) where
"*fn* F = Ifun (star_n F)"
- InternalFuns :: "('a star => 'b star) set"
+definition
+ InternalFuns :: "('a star => 'b star) set" where
"InternalFuns = {X. \<exists>F. X = *fn* F}"
--- a/src/HOL/Hyperreal/StarDef.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/StarDef.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,7 +13,7 @@
subsection {* A Free Ultrafilter over the Naturals *}
definition
- FreeUltrafilterNat :: "nat set set" ("\<U>")
+ FreeUltrafilterNat :: "nat set set" ("\<U>") where
"\<U> = (SOME U. freeultrafilter U)"
lemma freeultrafilter_FUFNat: "freeultrafilter \<U>"
@@ -36,14 +36,14 @@
subsection {* Definition of @{text star} type constructor *}
definition
- starrel :: "((nat \<Rightarrow> 'a) \<times> (nat \<Rightarrow> 'a)) set"
+ starrel :: "((nat \<Rightarrow> 'a) \<times> (nat \<Rightarrow> 'a)) set" where
"starrel = {(X,Y). {n. X n = Y n} \<in> \<U>}"
typedef 'a star = "(UNIV :: (nat \<Rightarrow> 'a) set) // starrel"
by (auto intro: quotientI)
definition
- star_n :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a star"
+ star_n :: "(nat \<Rightarrow> 'a) \<Rightarrow> 'a star" where
"star_n X = Abs_star (starrel `` {X})"
theorem star_cases [case_names star_n, cases type: star]:
@@ -157,10 +157,11 @@
subsection {* Standard elements *}
definition
- star_of :: "'a \<Rightarrow> 'a star"
+ star_of :: "'a \<Rightarrow> 'a star" where
"star_of x == star_n (\<lambda>n. x)"
- Standard :: "'a star set"
+definition
+ Standard :: "'a star set" where
"Standard = range star_of"
text {* Transfer tactic should remove occurrences of @{term star_of} *}
@@ -178,7 +179,7 @@
subsection {* Internal functions *}
definition
- Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" ("_ \<star> _" [300,301] 300)
+ Ifun :: "('a \<Rightarrow> 'b) star \<Rightarrow> 'a star \<Rightarrow> 'b star" ("_ \<star> _" [300,301] 300) where
"Ifun f \<equiv> \<lambda>x. Abs_star
(\<Union>F\<in>Rep_star f. \<Union>X\<in>Rep_star x. starrel``{\<lambda>n. F n (X n)})"
@@ -207,12 +208,12 @@
text {* Nonstandard extensions of functions *}
definition
- starfun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a star \<Rightarrow> 'b star)"
- ("*f* _" [80] 80)
+ starfun :: "('a \<Rightarrow> 'b) \<Rightarrow> ('a star \<Rightarrow> 'b star)" ("*f* _" [80] 80) where
"starfun f == \<lambda>x. star_of f \<star> x"
+definition
starfun2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a star \<Rightarrow> 'b star \<Rightarrow> 'c star)"
- ("*f2* _" [80] 80)
+ ("*f2* _" [80] 80) where
"starfun2 f == \<lambda>x y. star_of f \<star> x \<star> y"
declare starfun_def [transfer_unfold]
@@ -242,7 +243,7 @@
subsection {* Internal predicates *}
definition
- unstar :: "bool star \<Rightarrow> bool"
+ unstar :: "bool star \<Rightarrow> bool" where
"unstar b = (b = star_of True)"
lemma unstar_star_n: "unstar (star_n P) = ({n. P n} \<in> \<U>)"
@@ -259,12 +260,11 @@
by (simp only: unstar_star_n)
definition
- starP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> bool"
- ("*p* _" [80] 80)
+ starP :: "('a \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> bool" ("*p* _" [80] 80) where
"*p* P = (\<lambda>x. unstar (star_of P \<star> x))"
- starP2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> bool"
- ("*p2* _" [80] 80)
+definition
+ starP2 :: "('a \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> 'a star \<Rightarrow> 'b star \<Rightarrow> bool" ("*p2* _" [80] 80) where
"*p2* P = (\<lambda>x y. unstar (star_of P \<star> x \<star> y))"
declare starP_def [transfer_unfold]
@@ -287,7 +287,7 @@
subsection {* Internal sets *}
definition
- Iset :: "'a set star \<Rightarrow> 'a star set"
+ Iset :: "'a set star \<Rightarrow> 'a star set" where
"Iset A = {x. ( *p2* op \<in>) x A}"
lemma Iset_star_n:
@@ -329,7 +329,7 @@
text {* Nonstandard extensions of sets. *}
definition
- starset :: "'a set \<Rightarrow> 'a star set" ("*s* _" [80] 80)
+ starset :: "'a set \<Rightarrow> 'a star set" ("*s* _" [80] 80) where
"starset A = Iset (star_of A)"
declare starset_def [transfer_unfold]
--- a/src/HOL/Hyperreal/Transcendental.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Hyperreal/Transcendental.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,37 +12,45 @@
begin
definition
-
- exp :: "real => real"
+ exp :: "real => real" where
"exp x = (\<Sum>n. inverse(real (fact n)) * (x ^ n))"
- sin :: "real => real"
+definition
+ sin :: "real => real" where
"sin x = (\<Sum>n. (if even(n) then 0 else
((- 1) ^ ((n - Suc 0) div 2))/(real (fact n))) * x ^ n)"
- diffs :: "(nat => real) => nat => real"
+definition
+ diffs :: "(nat => real) => nat => real" where
"diffs c = (%n. real (Suc n) * c(Suc n))"
- cos :: "real => real"
+definition
+ cos :: "real => real" where
"cos x = (\<Sum>n. (if even(n) then ((- 1) ^ (n div 2))/(real (fact n))
else 0) * x ^ n)"
- ln :: "real => real"
+definition
+ ln :: "real => real" where
"ln x = (SOME u. exp u = x)"
- pi :: "real"
+definition
+ pi :: "real" where
"pi = 2 * (@x. 0 \<le> (x::real) & x \<le> 2 & cos x = 0)"
- tan :: "real => real"
+definition
+ tan :: "real => real" where
"tan x = (sin x)/(cos x)"
- arcsin :: "real => real"
+definition
+ arcsin :: "real => real" where
"arcsin y = (SOME x. -(pi/2) \<le> x & x \<le> pi/2 & sin x = y)"
- arcos :: "real => real"
+definition
+ arcos :: "real => real" where
"arcos y = (SOME x. 0 \<le> x & x \<le> pi & cos x = y)"
-
- arctan :: "real => real"
+
+definition
+ arctan :: "real => real" where
"arctan y = (SOME x. -(pi/2) < x & x < pi/2 & tan x = y)"
--- a/src/HOL/Induct/Comb.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/Comb.thy Fri Nov 17 02:20:03 2006 +0100
@@ -38,9 +38,11 @@
contract :: "(comb*comb) set"
abbreviation
- contract_rel1 :: "[comb,comb] => bool" (infixl "-1->" 50)
+ contract_rel1 :: "[comb,comb] => bool" (infixl "-1->" 50) where
"x -1-> y == (x,y) \<in> contract"
- contract_rel :: "[comb,comb] => bool" (infixl "--->" 50)
+
+abbreviation
+ contract_rel :: "[comb,comb] => bool" (infixl "--->" 50) where
"x ---> y == (x,y) \<in> contract^*"
inductive contract
@@ -59,9 +61,11 @@
parcontract :: "(comb*comb) set"
abbreviation
- parcontract_rel1 :: "[comb,comb] => bool" (infixl "=1=>" 50)
+ parcontract_rel1 :: "[comb,comb] => bool" (infixl "=1=>" 50) where
"x =1=> y == (x,y) \<in> parcontract"
- parcontract_rel :: "[comb,comb] => bool" (infixl "===>" 50)
+
+abbreviation
+ parcontract_rel :: "[comb,comb] => bool" (infixl "===>" 50) where
"x ===> y == (x,y) \<in> parcontract^*"
inductive parcontract
@@ -76,10 +80,11 @@
*}
definition
- I :: comb
+ I :: comb where
"I = S##K##K"
- diamond :: "('a * 'a)set => bool"
+definition
+ diamond :: "('a * 'a)set => bool" where
--{*confluence; Lambda/Commutation treats this more abstractly*}
"diamond(r) = (\<forall>x y. (x,y) \<in> r -->
(\<forall>y'. (x,y') \<in> r -->
--- a/src/HOL/Induct/LFilter.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/LFilter.thy Fri Nov 17 02:20:03 2006 +0100
@@ -20,10 +20,11 @@
declare findRel.intros [intro]
definition
- find :: "['a => bool, 'a llist] => 'a llist"
+ find :: "['a => bool, 'a llist] => 'a llist" where
"find p l = (SOME l'. (l,l'): findRel p | (l' = LNil & l ~: Domain(findRel p)))"
- lfilter :: "['a => bool, 'a llist] => 'a llist"
+definition
+ lfilter :: "['a => bool, 'a llist] => 'a llist" where
"lfilter p l = llist_corec l (%l. case find p l of
LNil => None
| LCons y z => Some(y,z))"
--- a/src/HOL/Induct/LList.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/LList.thy Fri Nov 17 02:20:03 2006 +0100
@@ -47,46 +47,54 @@
by (blast intro: llist.NIL_I)
definition
- list_Fun :: "['a item set, 'a item set] => 'a item set"
+ list_Fun :: "['a item set, 'a item set] => 'a item set" where
--{*Now used exclusively for abbreviating the coinduction rule*}
"list_Fun A X = {z. z = NIL | (\<exists>M a. z = CONS a M & a \<in> A & M \<in> X)}"
+definition
LListD_Fun ::
"[('a item * 'a item)set, ('a item * 'a item)set] =>
- ('a item * 'a item)set"
+ ('a item * 'a item)set" where
"LListD_Fun r X =
{z. z = (NIL, NIL) |
(\<exists>M N a b. z = (CONS a M, CONS b N) & (a, b) \<in> r & (M, N) \<in> X)}"
- LNil :: "'a llist"
+definition
+ LNil :: "'a llist" where
--{*abstract constructor*}
"LNil = Abs_LList NIL"
- LCons :: "['a, 'a llist] => 'a llist"
+definition
+ LCons :: "['a, 'a llist] => 'a llist" where
--{*abstract constructor*}
"LCons x xs = Abs_LList(CONS (Leaf x) (Rep_LList xs))"
- llist_case :: "['b, ['a, 'a llist]=>'b, 'a llist] => 'b"
+definition
+ llist_case :: "['b, ['a, 'a llist]=>'b, 'a llist] => 'b" where
"llist_case c d l =
List_case c (%x y. d (inv Leaf x) (Abs_LList y)) (Rep_LList l)"
- LList_corec_fun :: "[nat, 'a=> ('b item * 'a) option, 'a] => 'b item"
+definition
+ LList_corec_fun :: "[nat, 'a=> ('b item * 'a) option, 'a] => 'b item" where
"LList_corec_fun k f ==
nat_rec (%x. {})
(%j r x. case f x of None => NIL
| Some(z,w) => CONS z (r w))
k"
- LList_corec :: "['a, 'a => ('b item * 'a) option] => 'b item"
+definition
+ LList_corec :: "['a, 'a => ('b item * 'a) option] => 'b item" where
"LList_corec a f = (\<Union>k. LList_corec_fun k f a)"
- llist_corec :: "['a, 'a => ('b * 'a) option] => 'b llist"
+definition
+ llist_corec :: "['a, 'a => ('b * 'a) option] => 'b llist" where
"llist_corec a f =
Abs_LList(LList_corec a
(%z. case f z of None => None
| Some(v,w) => Some(Leaf(v), w)))"
- llistD_Fun :: "('a llist * 'a llist)set => ('a llist * 'a llist)set"
+definition
+ llistD_Fun :: "('a llist * 'a llist)set => ('a llist * 'a llist)set" where
"llistD_Fun(r) =
prod_fun Abs_LList Abs_LList `
LListD_Fun (diag(range Leaf))
@@ -105,25 +113,30 @@
subsubsection{* Sample function definitions. Item-based ones start with @{text L} *}
definition
- Lmap :: "('a item => 'b item) => ('a item => 'b item)"
+ Lmap :: "('a item => 'b item) => ('a item => 'b item)" where
"Lmap f M = LList_corec M (List_case None (%x M'. Some((f(x), M'))))"
- lmap :: "('a=>'b) => ('a llist => 'b llist)"
+definition
+ lmap :: "('a=>'b) => ('a llist => 'b llist)" where
"lmap f l = llist_corec l (%z. case z of LNil => None
| LCons y z => Some(f(y), z))"
- iterates :: "['a => 'a, 'a] => 'a llist"
+definition
+ iterates :: "['a => 'a, 'a] => 'a llist" where
"iterates f a = llist_corec a (%x. Some((x, f(x))))"
- Lconst :: "'a item => 'a item"
+definition
+ Lconst :: "'a item => 'a item" where
"Lconst(M) == lfp(%N. CONS M N)"
- Lappend :: "['a item, 'a item] => 'a item"
+definition
+ Lappend :: "['a item, 'a item] => 'a item" where
"Lappend M N = LList_corec (M,N)
(split(List_case (List_case None (%N1 N2. Some((N1, (NIL,N2)))))
(%M1 M2 N. Some((M1, (M2,N))))))"
- lappend :: "['a llist, 'a llist] => 'a llist"
+definition
+ lappend :: "['a llist, 'a llist] => 'a llist" where
"lappend l n = llist_corec (l,n)
(split(llist_case (llist_case None (%n1 n2. Some((n1, (LNil,n2)))))
(%l1 l2 n. Some((l1, (l2,n))))))"
--- a/src/HOL/Induct/Mutil.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/Mutil.thy Fri Nov 17 02:20:03 2006 +0100
@@ -31,13 +31,15 @@
text {* \medskip Sets of squares of the given colour*}
definition
- coloured :: "nat => (nat \<times> nat) set"
+ coloured :: "nat => (nat \<times> nat) set" where
"coloured b = {(i, j). (i + j) mod 2 = b}"
abbreviation
- whites :: "(nat \<times> nat) set"
+ whites :: "(nat \<times> nat) set" where
"whites == coloured 0"
- blacks :: "(nat \<times> nat) set"
+
+abbreviation
+ blacks :: "(nat \<times> nat) set" where
"blacks == coloured (Suc 0)"
--- a/src/HOL/Induct/Ordinals.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/Ordinals.thy Fri Nov 17 02:20:03 2006 +0100
@@ -32,9 +32,11 @@
"iter f (Suc n) = f \<circ> (iter f n)"
definition
- OpLim :: "(nat => (ordinal => ordinal)) => (ordinal => ordinal)"
+ OpLim :: "(nat => (ordinal => ordinal)) => (ordinal => ordinal)" where
"OpLim F a = Limit (\<lambda>n. F n a)"
- OpItw :: "(ordinal => ordinal) => (ordinal => ordinal)" ("\<Squnion>")
+
+definition
+ OpItw :: "(ordinal => ordinal) => (ordinal => ordinal)" ("\<Squnion>") where
"\<Squnion>f = OpLim (iter f)"
consts
@@ -52,7 +54,7 @@
"\<nabla>f (Limit h) = Limit (\<lambda>n. \<nabla>f (h n))"
definition
- deriv :: "(ordinal => ordinal) => (ordinal => ordinal)"
+ deriv :: "(ordinal => ordinal) => (ordinal => ordinal)" where
"deriv f = \<nabla>(\<Squnion>f)"
consts
@@ -62,9 +64,8 @@
"veblen (Succ a) = \<nabla>(OpLim (iter (veblen a)))"
"veblen (Limit f) = \<nabla>(OpLim (\<lambda>n. veblen (f n)))"
-definition
- "veb a = veblen a Zero"
- "\<epsilon>\<^isub>0 = veb Zero"
- "\<Gamma>\<^isub>0 = Limit (\<lambda>n. iter veb n Zero)"
+definition "veb a = veblen a Zero"
+definition "\<epsilon>\<^isub>0 = veb Zero"
+definition "\<Gamma>\<^isub>0 = Limit (\<lambda>n. iter veb n Zero)"
end
--- a/src/HOL/Induct/PropLog.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/PropLog.thy Fri Nov 17 02:20:03 2006 +0100
@@ -30,7 +30,7 @@
thms :: "'a pl set => 'a pl set"
abbreviation
- thm_rel :: "['a pl set, 'a pl] => bool" (infixl "|-" 50)
+ thm_rel :: "['a pl set, 'a pl] => bool" (infixl "|-" 50) where
"H |- p == p \<in> thms H"
inductive "thms(H)"
@@ -73,7 +73,7 @@
*}
definition
- sat :: "['a pl set, 'a pl] => bool" (infixl "|=" 50)
+ sat :: "['a pl set, 'a pl] => bool" (infixl "|=" 50) where
"H |= p = (\<forall>tt. (\<forall>q\<in>H. tt[[q]]) --> tt[[p]])"
--- a/src/HOL/Induct/QuoDataType.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/QuoDataType.thy Fri Nov 17 02:20:03 2006 +0100
@@ -23,7 +23,7 @@
consts msgrel :: "(freemsg * freemsg) set"
abbreviation
- msg_rel :: "[freemsg, freemsg] => bool" (infixl "~~" 50)
+ msg_rel :: "[freemsg, freemsg] => bool" (infixl "~~" 50) where
"X ~~ Y == (X,Y) \<in> msgrel"
notation (xsymbols)
@@ -143,18 +143,21 @@
text{*The abstract message constructors*}
definition
- Nonce :: "nat \<Rightarrow> msg"
+ Nonce :: "nat \<Rightarrow> msg" where
"Nonce N = Abs_Msg(msgrel``{NONCE N})"
- MPair :: "[msg,msg] \<Rightarrow> msg"
+definition
+ MPair :: "[msg,msg] \<Rightarrow> msg" where
"MPair X Y =
Abs_Msg (\<Union>U \<in> Rep_Msg X. \<Union>V \<in> Rep_Msg Y. msgrel``{MPAIR U V})"
- Crypt :: "[nat,msg] \<Rightarrow> msg"
+definition
+ Crypt :: "[nat,msg] \<Rightarrow> msg" where
"Crypt K X =
Abs_Msg (\<Union>U \<in> Rep_Msg X. msgrel``{CRYPT K U})"
- Decrypt :: "[nat,msg] \<Rightarrow> msg"
+definition
+ Decrypt :: "[nat,msg] \<Rightarrow> msg" where
"Decrypt K X =
Abs_Msg (\<Union>U \<in> Rep_Msg X. msgrel``{DECRYPT K U})"
@@ -228,7 +231,7 @@
subsection{*The Abstract Function to Return the Set of Nonces*}
definition
- nonces :: "msg \<Rightarrow> nat set"
+ nonces :: "msg \<Rightarrow> nat set" where
"nonces X = (\<Union>U \<in> Rep_Msg X. freenonces U)"
lemma nonces_congruent: "freenonces respects msgrel"
@@ -263,7 +266,7 @@
subsection{*The Abstract Function to Return the Left Part*}
definition
- left :: "msg \<Rightarrow> msg"
+ left :: "msg \<Rightarrow> msg" where
"left X = Abs_Msg (\<Union>U \<in> Rep_Msg X. msgrel``{freeleft U})"
lemma left_congruent: "(\<lambda>U. msgrel `` {freeleft U}) respects msgrel"
@@ -297,7 +300,7 @@
subsection{*The Abstract Function to Return the Right Part*}
definition
- right :: "msg \<Rightarrow> msg"
+ right :: "msg \<Rightarrow> msg" where
"right X = Abs_Msg (\<Union>U \<in> Rep_Msg X. msgrel``{freeright U})"
lemma right_congruent: "(\<lambda>U. msgrel `` {freeright U}) respects msgrel"
@@ -432,7 +435,7 @@
need this function in order to prove discrimination theorems.*}
definition
- discrim :: "msg \<Rightarrow> int"
+ discrim :: "msg \<Rightarrow> int" where
"discrim X = contents (\<Union>U \<in> Rep_Msg X. {freediscrim U})"
lemma discrim_congruent: "(\<lambda>U. {freediscrim U}) respects msgrel"
--- a/src/HOL/Induct/QuoNestedDataType.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/QuoNestedDataType.thy Fri Nov 17 02:20:03 2006 +0100
@@ -21,7 +21,7 @@
consts exprel :: "(freeExp * freeExp) set"
abbreviation
- exp_rel :: "[freeExp, freeExp] => bool" (infixl "~~" 50)
+ exp_rel :: "[freeExp, freeExp] => bool" (infixl "~~" 50) where
"X ~~ Y == (X,Y) \<in> exprel"
notation (xsymbols)
@@ -160,14 +160,16 @@
text{*The abstract message constructors*}
definition
- Var :: "nat \<Rightarrow> exp"
+ Var :: "nat \<Rightarrow> exp" where
"Var N = Abs_Exp(exprel``{VAR N})"
- Plus :: "[exp,exp] \<Rightarrow> exp"
+definition
+ Plus :: "[exp,exp] \<Rightarrow> exp" where
"Plus X Y =
Abs_Exp (\<Union>U \<in> Rep_Exp X. \<Union>V \<in> Rep_Exp Y. exprel``{PLUS U V})"
- FnCall :: "[nat, exp list] \<Rightarrow> exp"
+definition
+ FnCall :: "[nat, exp list] \<Rightarrow> exp" where
"FnCall F Xs =
Abs_Exp (\<Union>Us \<in> listset (map Rep_Exp Xs). exprel `` {FNCALL F Us})"
@@ -207,7 +209,7 @@
list of concrete expressions*}
definition
- Abs_ExpList :: "freeExp list => exp list"
+ Abs_ExpList :: "freeExp list => exp list" where
"Abs_ExpList Xs = map (%U. Abs_Exp(exprel``{U})) Xs"
lemma Abs_ExpList_Nil [simp]: "Abs_ExpList [] == []"
@@ -285,7 +287,7 @@
subsection{*The Abstract Function to Return the Set of Variables*}
definition
- vars :: "exp \<Rightarrow> nat set"
+ vars :: "exp \<Rightarrow> nat set" where
"vars X = (\<Union>U \<in> Rep_Exp X. freevars U)"
lemma vars_respects: "freevars respects exprel"
@@ -351,7 +353,7 @@
subsection{*Injectivity of @{term FnCall}*}
definition
- "fun" :: "exp \<Rightarrow> nat"
+ "fun" :: "exp \<Rightarrow> nat" where
"fun X = contents (\<Union>U \<in> Rep_Exp X. {freefun U})"
lemma fun_respects: "(%U. {freefun U}) respects exprel"
@@ -363,7 +365,7 @@
done
definition
- args :: "exp \<Rightarrow> exp list"
+ args :: "exp \<Rightarrow> exp list" where
"args X = contents (\<Union>U \<in> Rep_Exp X. {Abs_ExpList (freeargs U)})"
text{*This result can probably be generalized to arbitrary equivalence
@@ -398,7 +400,7 @@
function in order to prove discrimination theorems.*}
definition
- discrim :: "exp \<Rightarrow> int"
+ discrim :: "exp \<Rightarrow> int" where
"discrim X = contents (\<Union>U \<in> Rep_Exp X. {freediscrim U})"
lemma discrim_respects: "(\<lambda>U. {freediscrim U}) respects exprel"
--- a/src/HOL/Induct/SList.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/SList.thy Fri Nov 17 02:20:03 2006 +0100
@@ -37,10 +37,11 @@
(* Defining the Concrete Constructors *)
definition
- NIL :: "'a item"
+ NIL :: "'a item" where
"NIL = In0(Numb(0))"
- CONS :: "['a item, 'a item] => 'a item"
+definition
+ CONS :: "['a item, 'a item] => 'a item" where
"CONS M N = In1(Scons M N)"
consts
@@ -55,15 +56,15 @@
'a list = "list(range Leaf) :: 'a item set"
by (blast intro: list.NIL_I)
-abbreviation
- "Case == Datatype.Case"
- "Split == Datatype.Split"
+abbreviation "Case == Datatype.Case"
+abbreviation "Split == Datatype.Split"
definition
- List_case :: "['b, ['a item, 'a item]=>'b, 'a item] => 'b"
+ List_case :: "['b, ['a item, 'a item]=>'b, 'a item] => 'b" where
"List_case c d = Case(%x. c)(Split(d))"
- List_rec :: "['a item, 'b, ['a item, 'a item, 'b]=>'b] => 'b"
+definition
+ List_rec :: "['a item, 'b, ['a item, 'a item, 'b]=>'b] => 'b" where
"List_rec M c d = wfrec (trancl pred_sexp)
(%g. List_case c (%x y. d x y (g y))) M"
@@ -84,18 +85,21 @@
Cons :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" (infixr "#" 65)
definition
- Nil :: "'a list" ("[]")
+ Nil :: "'a list" ("[]") where
"Nil = Abs_List(NIL)"
- "Cons" :: "['a, 'a list] => 'a list" (infixr "#" 65)
+definition
+ "Cons" :: "['a, 'a list] => 'a list" (infixr "#" 65) where
"x#xs = Abs_List(CONS (Leaf x)(Rep_List xs))"
+definition
(* list Recursion -- the trancl is Essential; see list.ML *)
- list_rec :: "['a list, 'b, ['a, 'a list, 'b]=>'b] => 'b"
+ list_rec :: "['a list, 'b, ['a, 'a list, 'b]=>'b] => 'b" where
"list_rec l c d =
List_rec(Rep_List l) c (%x y r. d(inv Leaf x)(Abs_List y) r)"
- list_case :: "['b, ['a, 'a list]=>'b, 'a list] => 'b"
+definition
+ list_case :: "['b, ['a, 'a list]=>'b, 'a list] => 'b" where
"list_case a f xs = list_rec xs a (%x xs r. f x xs)"
(* list Enumeration *)
@@ -116,80 +120,98 @@
(* Generalized Map Functionals *)
definition
- Rep_map :: "('b => 'a item) => ('b list => 'a item)"
+ Rep_map :: "('b => 'a item) => ('b list => 'a item)" where
"Rep_map f xs = list_rec xs NIL(%x l r. CONS(f x) r)"
- Abs_map :: "('a item => 'b) => 'a item => 'b list"
+definition
+ Abs_map :: "('a item => 'b) => 'a item => 'b list" where
"Abs_map g M = List_rec M Nil (%N L r. g(N)#r)"
(**** Function definitions ****)
definition
-
- null :: "'a list => bool"
+ null :: "'a list => bool" where
"null xs = list_rec xs True (%x xs r. False)"
- hd :: "'a list => 'a"
+definition
+ hd :: "'a list => 'a" where
"hd xs = list_rec xs (@x. True) (%x xs r. x)"
- tl :: "'a list => 'a list"
+definition
+ tl :: "'a list => 'a list" where
"tl xs = list_rec xs (@xs. True) (%x xs r. xs)"
+definition
(* a total version of tl: *)
- ttl :: "'a list => 'a list"
+ ttl :: "'a list => 'a list" where
"ttl xs = list_rec xs [] (%x xs r. xs)"
- member :: "['a, 'a list] => bool" (infixl "mem" 55)
+definition
+ member :: "['a, 'a list] => bool" (infixl "mem" 55) where
"x mem xs = list_rec xs False (%y ys r. if y=x then True else r)"
- list_all :: "('a => bool) => ('a list => bool)"
+definition
+ list_all :: "('a => bool) => ('a list => bool)" where
"list_all P xs = list_rec xs True(%x l r. P(x) & r)"
- map :: "('a=>'b) => ('a list => 'b list)"
+definition
+ map :: "('a=>'b) => ('a list => 'b list)" where
"map f xs = list_rec xs [] (%x l r. f(x)#r)"
-
- append :: "['a list, 'a list] => 'a list" (infixr "@" 65)
+definition
+ append :: "['a list, 'a list] => 'a list" (infixr "@" 65) where
"xs@ys = list_rec xs ys (%x l r. x#r)"
- filter :: "['a => bool, 'a list] => 'a list"
+definition
+ filter :: "['a => bool, 'a list] => 'a list" where
"filter P xs = list_rec xs [] (%x xs r. if P(x)then x#r else r)"
- foldl :: "[['b,'a] => 'b, 'b, 'a list] => 'b"
+definition
+ foldl :: "[['b,'a] => 'b, 'b, 'a list] => 'b" where
"foldl f a xs = list_rec xs (%a. a)(%x xs r.%a. r(f a x))(a)"
- foldr :: "[['a,'b] => 'b, 'b, 'a list] => 'b"
+definition
+ foldr :: "[['a,'b] => 'b, 'b, 'a list] => 'b" where
"foldr f a xs = list_rec xs a (%x xs r. (f x r))"
- length :: "'a list => nat"
+definition
+ length :: "'a list => nat" where
"length xs = list_rec xs 0 (%x xs r. Suc r)"
- drop :: "['a list,nat] => 'a list"
+definition
+ drop :: "['a list,nat] => 'a list" where
"drop t n = (nat_rec(%x. x)(%m r xs. r(ttl xs)))(n)(t)"
- copy :: "['a, nat] => 'a list" (* make list of n copies of x *)
+definition
+ copy :: "['a, nat] => 'a list" where (* make list of n copies of x *)
"copy t = nat_rec [] (%m xs. t # xs)"
- flat :: "'a list list => 'a list"
+definition
+ flat :: "'a list list => 'a list" where
"flat = foldr (op @) []"
- nth :: "[nat, 'a list] => 'a"
+definition
+ nth :: "[nat, 'a list] => 'a" where
"nth = nat_rec hd (%m r xs. r(tl xs))"
- rev :: "'a list => 'a list"
+definition
+ rev :: "'a list => 'a list" where
"rev xs = list_rec xs [] (%x xs xsa. xsa @ [x])"
(* miscellaneous definitions *)
- zipWith :: "['a * 'b => 'c, 'a list * 'b list] => 'c list"
+definition
+ zipWith :: "['a * 'b => 'c, 'a list * 'b list] => 'c list" where
"zipWith f S = (list_rec (fst S) (%T.[])
(%x xs r. %T. if null T then []
else f(x,hd T) # r(tl T)))(snd(S))"
- zip :: "'a list * 'b list => ('a*'b) list"
+definition
+ zip :: "'a list * 'b list => ('a*'b) list" where
"zip = zipWith (%s. s)"
- unzip :: "('a*'b) list => ('a list * 'b list)"
+definition
+ unzip :: "('a*'b) list => ('a list * 'b list)" where
"unzip = foldr(% (a,b)(c,d).(a#c,b#d))([],[])"
--- a/src/HOL/Induct/Sexp.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/Sexp.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,9 +11,8 @@
types
'a item = "'a Datatype.item"
-abbreviation
- "Leaf == Datatype.Leaf"
- "Numb == Datatype.Numb"
+abbreviation "Leaf == Datatype.Leaf"
+abbreviation "Numb == Datatype.Numb"
consts
sexp :: "'a item set"
@@ -26,16 +25,18 @@
definition
sexp_case :: "['a=>'b, nat=>'b, ['a item, 'a item]=>'b,
- 'a item] => 'b"
+ 'a item] => 'b" where
"sexp_case c d e M = (THE z. (EX x. M=Leaf(x) & z=c(x))
| (EX k. M=Numb(k) & z=d(k))
| (EX N1 N2. M = Scons N1 N2 & z=e N1 N2))"
- pred_sexp :: "('a item * 'a item)set"
+definition
+ pred_sexp :: "('a item * 'a item)set" where
"pred_sexp = (\<Union>M \<in> sexp. \<Union>N \<in> sexp. {(M, Scons M N), (N, Scons M N)})"
+definition
sexp_rec :: "['a item, 'a=>'b, nat=>'b,
- ['a item, 'a item, 'b, 'b]=>'b] => 'b"
+ ['a item, 'a item, 'b, 'b]=>'b] => 'b" where
"sexp_rec M c d e = wfrec pred_sexp
(%g. sexp_case c d (%N1 N2. e N1 N2 (g N1) (g N2))) M"
--- a/src/HOL/Induct/Tree.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Induct/Tree.thy Fri Nov 17 02:20:03 2006 +0100
@@ -72,10 +72,11 @@
closure. *}
definition
- brouwer_pred :: "(brouwer * brouwer) set"
+ brouwer_pred :: "(brouwer * brouwer) set" where
"brouwer_pred = (\<Union>i. {(m,n). n = Succ m \<or> (EX f. n = Lim f & m = f i)})"
- brouwer_order :: "(brouwer * brouwer) set"
+definition
+ brouwer_order :: "(brouwer * brouwer) set" where
"brouwer_order = brouwer_pred^+"
lemma wf_brouwer_pred: "wf brouwer_pred"
--- a/src/HOL/Integ/IntDef.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Integ/IntDef.thy Fri Nov 17 02:20:03 2006 +0100
@@ -870,9 +870,11 @@
*}
definition
- int_aux :: "int \<Rightarrow> nat \<Rightarrow> int"
+ int_aux :: "int \<Rightarrow> nat \<Rightarrow> int" where
"int_aux i n = (i + int n)"
- nat_aux :: "nat \<Rightarrow> int \<Rightarrow> nat"
+
+definition
+ nat_aux :: "nat \<Rightarrow> int \<Rightarrow> nat" where
"nat_aux n i = (n + nat i)"
lemma [code]:
--- a/src/HOL/Integ/NatBin.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Integ/NatBin.thy Fri Nov 17 02:20:03 2006 +0100
@@ -20,7 +20,7 @@
nat_number_of_def: "number_of v == nat (number_of (v\<Colon>int))"
abbreviation (xsymbols)
- square :: "'a::power => 'a" ("(_\<twosuperior>)" [1000] 999)
+ square :: "'a::power => 'a" ("(_\<twosuperior>)" [1000] 999) where
"x\<twosuperior> == x^2"
notation (latex output)
--- a/src/HOL/Integ/Numeral.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Integ/Numeral.thy Fri Nov 17 02:20:03 2006 +0100
@@ -51,6 +51,8 @@
abbreviation
"Numeral0 \<equiv> number_of Pls"
+
+abbreviation
"Numeral1 \<equiv> number_of (Pls BIT B1)"
lemma Let_number_of [simp]: "Let (number_of v) f = f (number_of v)"
@@ -64,9 +66,11 @@
unfolding Let_def ..
definition
- succ :: "int \<Rightarrow> int"
+ succ :: "int \<Rightarrow> int" where
"succ k = k + 1"
- pred :: "int \<Rightarrow> int"
+
+definition
+ pred :: "int \<Rightarrow> int" where
"pred k = k - 1"
lemmas numeral_simps =
--- a/src/HOL/Isar_examples/Hoare.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Isar_examples/Hoare.thy Fri Nov 17 02:20:03 2006 +0100
@@ -32,7 +32,7 @@
| While "'a bexp" "'a assn" "'a com"
abbreviation
- Skip ("SKIP")
+ Skip ("SKIP") where
"SKIP == Basic id"
types
--- a/src/HOL/Lambda/Commutation.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/Commutation.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,22 +11,25 @@
subsection {* Basic definitions *}
definition
- square :: "[('a \<times> 'a) set, ('a \<times> 'a) set, ('a \<times> 'a) set, ('a \<times> 'a) set] => bool"
+ square :: "[('a \<times> 'a) set, ('a \<times> 'a) set, ('a \<times> 'a) set, ('a \<times> 'a) set] => bool" where
"square R S T U =
(\<forall>x y. (x, y) \<in> R --> (\<forall>z. (x, z) \<in> S --> (\<exists>u. (y, u) \<in> T \<and> (z, u) \<in> U)))"
- commute :: "[('a \<times> 'a) set, ('a \<times> 'a) set] => bool"
+definition
+ commute :: "[('a \<times> 'a) set, ('a \<times> 'a) set] => bool" where
"commute R S = square R S S R"
- diamond :: "('a \<times> 'a) set => bool"
+definition
+ diamond :: "('a \<times> 'a) set => bool" where
"diamond R = commute R R"
- Church_Rosser :: "('a \<times> 'a) set => bool"
+definition
+ Church_Rosser :: "('a \<times> 'a) set => bool" where
"Church_Rosser R =
(\<forall>x y. (x, y) \<in> (R \<union> R^-1)^* --> (\<exists>z. (x, z) \<in> R^* \<and> (y, z) \<in> R^*))"
abbreviation
- confluent :: "('a \<times> 'a) set => bool"
+ confluent :: "('a \<times> 'a) set => bool" where
"confluent R == diamond (R^*)"
--- a/src/HOL/Lambda/Eta.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/Eta.thy Fri Nov 17 02:20:03 2006 +0100
@@ -22,11 +22,15 @@
eta :: "(dB \<times> dB) set"
abbreviation
- eta_red :: "[dB, dB] => bool" (infixl "-e>" 50)
+ eta_red :: "[dB, dB] => bool" (infixl "-e>" 50) where
"s -e> t == (s, t) \<in> eta"
- eta_reds :: "[dB, dB] => bool" (infixl "-e>>" 50)
+
+abbreviation
+ eta_reds :: "[dB, dB] => bool" (infixl "-e>>" 50) where
"s -e>> t == (s, t) \<in> eta^*"
- eta_red0 :: "[dB, dB] => bool" (infixl "-e>=" 50)
+
+abbreviation
+ eta_red0 :: "[dB, dB] => bool" (infixl "-e>=" 50) where
"s -e>= t == (s, t) \<in> eta^="
inductive eta
@@ -224,7 +228,7 @@
par_eta :: "(dB \<times> dB) set"
abbreviation
- par_eta_red :: "[dB, dB] => bool" (infixl "=e>" 50)
+ par_eta_red :: "[dB, dB] => bool" (infixl "=e>" 50) where
"s =e> t == (s, t) \<in> par_eta"
notation (xsymbols)
--- a/src/HOL/Lambda/Lambda.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/Lambda.thy Fri Nov 17 02:20:03 2006 +0100
@@ -57,13 +57,15 @@
beta :: "(dB \<times> dB) set"
abbreviation
- beta_red :: "[dB, dB] => bool" (infixl "->" 50)
+ beta_red :: "[dB, dB] => bool" (infixl "->" 50) where
"s -> t == (s, t) \<in> beta"
- beta_reds :: "[dB, dB] => bool" (infixl "->>" 50)
+
+abbreviation
+ beta_reds :: "[dB, dB] => bool" (infixl "->>" 50) where
"s ->> t == (s, t) \<in> beta^*"
notation (latex)
- beta_red (infixl "\<rightarrow>\<^sub>\<beta>" 50)
+ beta_red (infixl "\<rightarrow>\<^sub>\<beta>" 50) and
beta_reds (infixl "\<rightarrow>\<^sub>\<beta>\<^sup>*" 50)
inductive beta
--- a/src/HOL/Lambda/ListApplication.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/ListApplication.thy Fri Nov 17 02:20:03 2006 +0100
@@ -9,7 +9,7 @@
theory ListApplication imports Lambda begin
abbreviation
- list_application :: "dB => dB list => dB" (infixl "\<degree>\<degree>" 150)
+ list_application :: "dB => dB list => dB" (infixl "\<degree>\<degree>" 150) where
"t \<degree>\<degree> ts == foldl (op \<degree>) t ts"
lemma apps_eq_tail_conv [iff]: "(r \<degree>\<degree> ts = s \<degree>\<degree> ts) = (r = s)"
--- a/src/HOL/Lambda/ListBeta.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/ListBeta.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,7 +13,7 @@
*}
abbreviation
- list_beta :: "dB list => dB list => bool" (infixl "=>" 50)
+ list_beta :: "dB list => dB list => bool" (infixl "=>" 50) where
"rs => ss == (rs, ss) : step1 beta"
lemma head_Var_reduction:
--- a/src/HOL/Lambda/ListOrder.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/ListOrder.thy Fri Nov 17 02:20:03 2006 +0100
@@ -14,7 +14,7 @@
*}
definition
- step1 :: "('a \<times> 'a) set => ('a list \<times> 'a list) set"
+ step1 :: "('a \<times> 'a) set => ('a list \<times> 'a list) set" where
"step1 r =
{(ys, xs). \<exists>us z z' vs. xs = us @ z # vs \<and> (z', z) \<in> r \<and> ys =
us @ z' # vs}"
--- a/src/HOL/Lambda/ParRed.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/ParRed.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,7 +18,7 @@
par_beta :: "(dB \<times> dB) set"
abbreviation
- par_beta_red :: "[dB, dB] => bool" (infixl "=>" 50)
+ par_beta_red :: "[dB, dB] => bool" (infixl "=>" 50) where
"s => t == (s, t) \<in> par_beta"
inductive par_beta
--- a/src/HOL/Lambda/StrongNorm.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/StrongNorm.thy Fri Nov 17 02:20:03 2006 +0100
@@ -190,7 +190,7 @@
by (rule typing.App)
qed
with Cons True show ?thesis
- by (simp add: map_compose [symmetric] o_def)
+ by (simp add: map_compose [symmetric] comp_def)
qed
next
case False
--- a/src/HOL/Lambda/Type.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/Type.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,7 +12,7 @@
subsection {* Environments *}
definition
- shift :: "(nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a" ("_<_:_>" [90, 0, 0] 91)
+ shift :: "(nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a" ("_<_:_>" [90, 0, 0] 91) where
"e<i:a> = (\<lambda>j. if j < i then e j else if j = i then a else e (j - 1))"
notation (xsymbols)
@@ -50,21 +50,23 @@
typings :: "(nat \<Rightarrow> type) \<Rightarrow> dB list \<Rightarrow> type list \<Rightarrow> bool"
abbreviation
- funs :: "type list \<Rightarrow> type \<Rightarrow> type" (infixr "=>>" 200)
+ funs :: "type list \<Rightarrow> type \<Rightarrow> type" (infixr "=>>" 200) where
"Ts =>> T == foldr Fun Ts T"
- typing_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ |- _ : _" [50, 50, 50] 50)
+abbreviation
+ typing_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ |- _ : _" [50, 50, 50] 50) where
"env |- t : T == (env, t, T) \<in> typing"
+abbreviation
typings_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB list \<Rightarrow> type list \<Rightarrow> bool"
- ("_ ||- _ : _" [50, 50, 50] 50)
+ ("_ ||- _ : _" [50, 50, 50] 50) where
"env ||- ts : Ts == typings env ts Ts"
notation (xsymbols)
typing_rel ("_ \<turnstile> _ : _" [50, 50, 50] 50)
notation (latex)
- funs (infixr "\<Rrightarrow>" 200)
+ funs (infixr "\<Rrightarrow>" 200) and
typings_rel ("_ \<tturnstile> _ : _" [50, 50, 50] 50)
inductive typing
--- a/src/HOL/Lambda/WeakNorm.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lambda/WeakNorm.thy Fri Nov 17 02:20:03 2006 +0100
@@ -17,7 +17,7 @@
subsection {* Terms in normal form *}
definition
- listall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
+ listall :: "('a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool" where
"listall P xs == (\<forall>i. i < length xs \<longrightarrow> P (xs ! i))"
declare listall_def [extraction_expand]
@@ -362,7 +362,7 @@
rtyping :: "((nat \<Rightarrow> type) \<times> dB \<times> type) set"
abbreviation
- rtyping_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ |-\<^sub>R _ : _" [50, 50, 50] 50)
+ rtyping_rel :: "(nat \<Rightarrow> type) \<Rightarrow> dB \<Rightarrow> type \<Rightarrow> bool" ("_ |-\<^sub>R _ : _" [50, 50, 50] 50) where
"e |-\<^sub>R t : T == (e, t, T) \<in> rtyping"
notation (xsymbols)
--- a/src/HOL/Lattice/Bounds.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lattice/Bounds.thy Fri Nov 17 02:20:03 2006 +0100
@@ -16,16 +16,19 @@
*}
definition
- is_inf :: "'a::partial_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+ is_inf :: "'a::partial_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
"is_inf x y inf = (inf \<sqsubseteq> x \<and> inf \<sqsubseteq> y \<and> (\<forall>z. z \<sqsubseteq> x \<and> z \<sqsubseteq> y \<longrightarrow> z \<sqsubseteq> inf))"
- is_sup :: "'a::partial_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+definition
+ is_sup :: "'a::partial_order \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool" where
"is_sup x y sup = (x \<sqsubseteq> sup \<and> y \<sqsubseteq> sup \<and> (\<forall>z. x \<sqsubseteq> z \<and> y \<sqsubseteq> z \<longrightarrow> sup \<sqsubseteq> z))"
- is_Inf :: "'a::partial_order set \<Rightarrow> 'a \<Rightarrow> bool"
+definition
+ is_Inf :: "'a::partial_order set \<Rightarrow> 'a \<Rightarrow> bool" where
"is_Inf A inf = ((\<forall>x \<in> A. inf \<sqsubseteq> x) \<and> (\<forall>z. (\<forall>x \<in> A. z \<sqsubseteq> x) \<longrightarrow> z \<sqsubseteq> inf))"
- is_Sup :: "'a::partial_order set \<Rightarrow> 'a \<Rightarrow> bool"
+definition
+ is_Sup :: "'a::partial_order set \<Rightarrow> 'a \<Rightarrow> bool" where
"is_Sup A sup = ((\<forall>x \<in> A. x \<sqsubseteq> sup) \<and> (\<forall>z. (\<forall>x \<in> A. x \<sqsubseteq> z) \<longrightarrow> sup \<sqsubseteq> z))"
text {*
--- a/src/HOL/Lattice/CompleteLattice.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lattice/CompleteLattice.thy Fri Nov 17 02:20:03 2006 +0100
@@ -32,13 +32,14 @@
*}
definition
- Meet :: "'a::complete_lattice set \<Rightarrow> 'a"
+ Meet :: "'a::complete_lattice set \<Rightarrow> 'a" where
"Meet A = (THE inf. is_Inf A inf)"
- Join :: "'a::complete_lattice set \<Rightarrow> 'a"
+definition
+ Join :: "'a::complete_lattice set \<Rightarrow> 'a" where
"Join A = (THE sup. is_Sup A sup)"
notation (xsymbols)
- Meet ("\<Sqinter>_" [90] 90)
+ Meet ("\<Sqinter>_" [90] 90) and
Join ("\<Squnion>_" [90] 90)
text {*
@@ -143,9 +144,10 @@
*}
definition
- bottom :: "'a::complete_lattice" ("\<bottom>")
+ bottom :: "'a::complete_lattice" ("\<bottom>") where
"\<bottom> = \<Sqinter>UNIV"
- top :: "'a::complete_lattice" ("\<top>")
+definition
+ top :: "'a::complete_lattice" ("\<top>") where
"\<top> = \<Squnion>UNIV"
lemma bottom_least [intro?]: "\<bottom> \<sqsubseteq> x"
--- a/src/HOL/Lattice/Lattice.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Lattice/Lattice.thy Fri Nov 17 02:20:03 2006 +0100
@@ -25,13 +25,14 @@
*}
definition
- meet :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "&&" 70)
+ meet :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "&&" 70) where
"x && y = (THE inf. is_inf x y inf)"
- join :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "||" 65)
+definition
+ join :: "'a::lattice \<Rightarrow> 'a \<Rightarrow> 'a" (infixl "||" 65) where
"x || y = (THE sup. is_sup x y sup)"
notation (xsymbols)
- meet (infixl "\<sqinter>" 70)
+ meet (infixl "\<sqinter>" 70) and
join (infixl "\<squnion>" 65)
text {*
@@ -337,9 +338,10 @@
*}
definition
- minimum :: "'a::linear_order \<Rightarrow> 'a \<Rightarrow> 'a"
+ minimum :: "'a::linear_order \<Rightarrow> 'a \<Rightarrow> 'a" where
"minimum x y = (if x \<sqsubseteq> y then x else y)"
- maximum :: "'a::linear_order \<Rightarrow> 'a \<Rightarrow> 'a"
+definition
+ maximum :: "'a::linear_order \<Rightarrow> 'a \<Rightarrow> 'a" where
"maximum x y = (if x \<sqsubseteq> y then y else x)"
lemma is_inf_minimum: "is_inf x y (minimum x y)"
--- a/src/HOL/Library/AssocList.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/AssocList.thy Fri Nov 17 02:20:03 2006 +0100
@@ -96,12 +96,12 @@
(* ******************************************************************************** *)
lemma delete_simps [simp]:
-"delete k [] = []"
-"delete k (p#ps) = (if fst p = k then delete k ps else p # delete k ps)"
+ "delete k [] = []"
+ "delete k (p#ps) = (if fst p = k then delete k ps else p # delete k ps)"
by (simp_all add: delete_def)
lemma delete_id[simp]: "k \<notin> fst ` set al \<Longrightarrow> delete k al = al"
-by(induct al, auto)
+ by (induct al) auto
lemma delete_conv: "map_of (delete k al) k' = ((map_of al)(k := None)) k'"
by (induct al) auto
@@ -112,9 +112,9 @@
lemma delete_idem: "delete k (delete k al) = delete k al"
by (induct al) auto
-lemma map_of_delete[simp]:
- "k' \<noteq> k \<Longrightarrow> map_of (delete k al) k' = map_of al k'"
-by(induct al, auto)
+lemma map_of_delete [simp]:
+ "k' \<noteq> k \<Longrightarrow> map_of (delete k al) k' = map_of al k'"
+ by (induct al) auto
lemma delete_notin_dom: "k \<notin> fst ` set (delete k al)"
by (induct al) auto
@@ -547,7 +547,7 @@
lemma compose_conv:
shows "map_of (compose xs ys) k = (map_of ys \<circ>\<^sub>m map_of xs) k"
-proof (induct xs ys rule: compose_induct )
+proof (induct xs ys rule: compose_induct)
case Nil thus ?case by simp
next
case (Cons x xs)
@@ -595,7 +595,7 @@
using prems by (simp add: compose_conv)
lemma dom_compose: "fst ` set (compose xs ys) \<subseteq> fst ` set xs"
-proof (induct xs ys rule: compose_induct )
+proof (induct xs ys rule: compose_induct)
case Nil thus ?case by simp
next
case (Cons x xs)
@@ -799,12 +799,12 @@
"map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
by (simp add: update_conv' restr_conv') (rule fun_upd_restrict)
-lemma upate_restr_conv[simp]:
+lemma upate_restr_conv [simp]:
"x \<in> D \<Longrightarrow>
map_of (update x y (restrict D al)) = map_of (update x y (restrict (D-{x}) al))"
by (simp add: update_conv' restr_conv')
-lemma restr_updates[simp]: "
+lemma restr_updates [simp]: "
\<lbrakk> length xs = length ys; set xs \<subseteq> D \<rbrakk>
\<Longrightarrow> map_of (restrict D (updates xs ys al)) =
map_of (updates xs ys (restrict (D - set xs) al))"
--- a/src/HOL/Library/BigO.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/BigO.thy Fri Nov 17 02:20:03 2006 +0100
@@ -39,7 +39,7 @@
subsection {* Definitions *}
definition
- bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set" ("(1O'(_'))")
+ bigo :: "('a => 'b::ordered_idom) => ('a => 'b) set" ("(1O'(_'))") where
"O(f::('a => 'b)) =
{h. EX c. ALL x. abs (h x) <= c * abs (f x)}"
@@ -736,7 +736,7 @@
definition
lesso :: "('a => 'b::ordered_idom) => ('a => 'b) => ('a => 'b)"
- (infixl "<o" 70)
+ (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/Library/Char_ord.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Char_ord.thy Fri Nov 17 02:20:03 2006 +0100
@@ -32,8 +32,8 @@
"nibble_to_int NibbleF = 15"
definition
- int_to_nibble :: "int \<Rightarrow> nibble"
- "int_to_nibble x \<equiv> (let y = x mod 16 in
+ int_to_nibble :: "int \<Rightarrow> nibble" where
+ "int_to_nibble x = (let y = x mod 16 in
if y = 0 then Nibble0 else
if y = 1 then Nibble1 else
if y = 2 then Nibble2 else
--- a/src/HOL/Library/Coinductive_List.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Coinductive_List.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,6 +13,7 @@
definition
"NIL = Datatype.In0 (Datatype.Numb 0)"
+definition
"CONS M N = Datatype.In1 (Datatype.Scons M N)"
lemma CONS_not_NIL [iff]: "CONS M N \<noteq> NIL"
@@ -146,6 +147,7 @@
definition
"LNil = Abs_llist NIL"
+definition
"LCons x xs = Abs_llist (CONS (Datatype.Leaf x) (Rep_llist xs))"
lemma LCons_not_LNil [iff]: "LCons x xs \<noteq> LNil"
@@ -573,6 +575,7 @@
definition
"Lmap f M = LList_corec M (List_case None (\<lambda>x M'. Some (f x, M')))"
+definition
"lmap f l = llist_corec l
(\<lambda>z.
case z of LNil \<Rightarrow> None
@@ -671,6 +674,7 @@
(split (List_case
(List_case None (\<lambda>N1 N2. Some (N1, (NIL, N2))))
(\<lambda>M1 M2 N. Some (M1, (M2, N)))))"
+definition
"lappend l n = llist_corec (l, n)
(split (llist_case
(llist_case None (\<lambda>n1 n2. Some (n1, (LNil, n2))))
@@ -746,7 +750,7 @@
text {* @{text llist_fun_equalityI} cannot be used here! *}
definition
- iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist"
+ iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a llist" where
"iterates f a = llist_corec a (\<lambda>x. Some (x, f x))"
lemma iterates: "iterates f x = LCons x (iterates f (f x))"
--- a/src/HOL/Library/Commutative_Ring.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Commutative_Ring.thy Fri Nov 17 02:20:03 2006 +0100
@@ -48,14 +48,14 @@
text {* Create polynomial normalized polynomials given normalized inputs. *}
definition
- mkPinj :: "nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
+ mkPinj :: "nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol" where
"mkPinj x P = (case P of
Pc c \<Rightarrow> Pc c |
Pinj y P \<Rightarrow> Pinj (x + y) P |
PX p1 y p2 \<Rightarrow> Pinj x P)"
definition
- mkPX :: "'a::{comm_ring,recpower} pol \<Rightarrow> nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
+ mkPX :: "'a::{comm_ring,recpower} pol \<Rightarrow> nat \<Rightarrow> 'a pol \<Rightarrow> 'a pol" where
"mkPX P i Q = (case P of
Pc c \<Rightarrow> (if (c = 0) then (mkPinj 1 Q) else (PX P i Q)) |
Pinj j R \<Rightarrow> PX P i Q |
@@ -128,7 +128,7 @@
text {* Substraction *}
definition
- sub :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol"
+ sub :: "'a::{comm_ring,recpower} pol \<Rightarrow> 'a pol \<Rightarrow> 'a pol" where
"sub p q = add (p, neg q)"
text {* Square for Fast Exponentation *}
--- a/src/HOL/Library/Continuity.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Continuity.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,8 +18,8 @@
"continuous F == !M. chain M \<longrightarrow> F(SUP i. M i) = (SUP i. F(M i))"
abbreviation
- bot :: "'a::order"
-"bot == Sup {}"
+ bot :: "'a::order" where
+ "bot == Sup {}"
lemma SUP_nat_conv:
"(SUP n::nat. M n::'a::comp_lat) = join (M 0) (SUP n. M(Suc n))"
@@ -91,7 +91,7 @@
subsection "Chains"
definition
- up_chain :: "(nat => 'a set) => bool"
+ up_chain :: "(nat => 'a set) => bool" where
"up_chain F = (\<forall>i. F i \<subseteq> F (Suc i))"
lemma up_chainI: "(!!i. F i \<subseteq> F (Suc i)) ==> up_chain F"
@@ -113,7 +113,7 @@
definition
- down_chain :: "(nat => 'a set) => bool"
+ down_chain :: "(nat => 'a set) => bool" where
"down_chain F = (\<forall>i. F (Suc i) \<subseteq> F i)"
lemma down_chainI: "(!!i. F (Suc i) \<subseteq> F i) ==> down_chain F"
@@ -137,7 +137,7 @@
subsection "Continuity"
definition
- up_cont :: "('a set => 'a set) => bool"
+ up_cont :: "('a set => 'a set) => bool" where
"up_cont f = (\<forall>F. up_chain F --> f (\<Union>(range F)) = \<Union>(f ` range F))"
lemma up_contI:
@@ -164,7 +164,7 @@
definition
- down_cont :: "('a set => 'a set) => bool"
+ down_cont :: "('a set => 'a set) => bool" where
"down_cont f =
(\<forall>F. down_chain F --> f (Inter (range F)) = Inter (f ` range F))"
@@ -194,7 +194,7 @@
subsection "Iteration"
definition
- up_iterate :: "('a set => 'a set) => nat => 'a set"
+ up_iterate :: "('a set => 'a set) => nat => 'a set" where
"up_iterate f n = (f^n) {}"
lemma up_iterate_0 [simp]: "up_iterate f 0 = {}"
@@ -246,7 +246,7 @@
definition
- down_iterate :: "('a set => 'a set) => nat => 'a set"
+ down_iterate :: "('a set => 'a set) => nat => 'a set" where
"down_iterate f n = (f^n) UNIV"
lemma down_iterate_0 [simp]: "down_iterate f 0 = UNIV"
--- a/src/HOL/Library/EfficientNat.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/EfficientNat.thy Fri Nov 17 02:20:03 2006 +0100
@@ -25,7 +25,7 @@
*}
definition
- nat_of_int :: "int \<Rightarrow> nat"
+ nat_of_int :: "int \<Rightarrow> nat" where
"k \<ge> 0 \<Longrightarrow> nat_of_int k = nat k"
lemma nat_of_int_int:
--- a/src/HOL/Library/ExecutableRat.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/ExecutableRat.thy Fri Nov 17 02:20:03 2006 +0100
@@ -27,8 +27,8 @@
instance erat :: ord ..
definition
- norm :: "erat \<Rightarrow> erat"
- norm_def: "norm r = (case r of (Rat a p q) \<Rightarrow>
+ norm :: "erat \<Rightarrow> erat" where
+ "norm r = (case r of (Rat a p q) \<Rightarrow>
if p = 0 then Rat True 0 1
else
let
@@ -36,19 +36,28 @@
in let
m = zgcd (absp, q)
in Rat (a = (0 <= p)) (absp div m) (q div m))"
- common :: "(int * int) * (int * int) \<Rightarrow> (int * int) * int"
- common_def: "common r = (case r of ((p1, q1), (p2, q2)) \<Rightarrow>
+
+definition
+ common :: "(int * int) * (int * int) \<Rightarrow> (int * int) * int" where
+ "common r = (case r of ((p1, q1), (p2, q2)) \<Rightarrow>
let q' = q1 * q2 div int (gcd (nat q1, nat q2))
in ((p1 * (q' div q1), p2 * (q' div q2)), q'))"
- of_quotient :: "int \<Rightarrow> int \<Rightarrow> erat"
- of_quotient_def: "of_quotient a b =
- norm (Rat True a b)"
- of_rat :: "rat \<Rightarrow> erat"
- of_rat_def: "of_rat r = split of_quotient (SOME s. s : Rep_Rat r)"
- to_rat :: "erat \<Rightarrow> rat"
- to_rat_def: "to_rat r = (case r of (Rat a p q) \<Rightarrow>
+
+definition
+ of_quotient :: "int \<Rightarrow> int \<Rightarrow> erat" where
+ "of_quotient a b = norm (Rat True a b)"
+
+definition
+ of_rat :: "rat \<Rightarrow> erat" where
+ "of_rat r = split of_quotient (SOME s. s : Rep_Rat r)"
+
+definition
+ to_rat :: "erat \<Rightarrow> rat" where
+ "to_rat r = (case r of (Rat a p q) \<Rightarrow>
if a then Fract p q else Fract (uminus p) q)"
- eq_erat :: "erat \<Rightarrow> erat \<Rightarrow> bool"
+
+definition
+ eq_erat :: "erat \<Rightarrow> erat \<Rightarrow> bool" where
"eq_erat r s = (norm r = norm s)"
axiomatization
--- a/src/HOL/Library/ExecutableSet.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/ExecutableSet.thy Fri Nov 17 02:20:03 2006 +0100
@@ -14,7 +14,7 @@
instance set :: (eq) eq ..
definition
- minus_set :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set"
+ minus_set :: "'a set \<Rightarrow> 'a set \<Rightarrow> 'a set" where
"minus_set xs ys = ys - xs"
lemma [code inline]:
@@ -22,8 +22,8 @@
unfolding minus_set_def ..
definition
- subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"
- subset_def: "subset = op \<subseteq>"
+ subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
+ "subset = op \<subseteq>"
lemmas [symmetric, code inline] = subset_def
@@ -44,8 +44,8 @@
unfolding bex_triv_one_point1 ..
definition
- filter_set :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set"
- filter_set_def: "filter_set P xs = {x\<in>xs. P x}"
+ filter_set :: "('a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> 'a set" where
+ "filter_set P xs = {x\<in>xs. P x}"
lemmas [symmetric, code inline] = filter_set_def
@@ -55,11 +55,15 @@
subsection {* Basic definitions *}
definition
- flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c"
+ flip :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'a \<Rightarrow> 'c" where
"flip f a b = f b a"
- member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool"
+
+definition
+ member :: "'a list \<Rightarrow> 'a \<Rightarrow> bool" where
"member xs x = (x \<in> set xs)"
- insertl :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
+
+definition
+ insertl :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"insertl x xs = (if member xs x then xs else x#xs)"
lemma
@@ -174,9 +178,11 @@
"intersects (x#xs) = foldr intersect xs x"
definition
- map_union :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list"
+ map_union :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
"map_union xs f = unions (map f xs)"
- map_inter :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list"
+
+definition
+ map_inter :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b list) \<Rightarrow> 'b list" where
"map_inter xs f = intersects (map f xs)"
@@ -237,9 +243,10 @@
unfolding map_inter_def iso_intersects by (induct xs) (simp_all add: iso_member, auto)
definition
- Blall :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+ Blall :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
"Blall = flip list_all"
- Blex :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
+definition
+ Blex :: "'a list \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
"Blex = flip list_ex"
lemma iso_Ball:
--- a/src/HOL/Library/FuncSet.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/FuncSet.thy Fri Nov 17 02:20:03 2006 +0100
@@ -10,17 +10,20 @@
begin
definition
- Pi :: "['a set, 'a => 'b set] => ('a => 'b) set"
+ Pi :: "['a set, 'a => 'b set] => ('a => 'b) set" where
"Pi A B = {f. \<forall>x. x \<in> A --> f x \<in> B x}"
- extensional :: "'a set => ('a => 'b) set"
+definition
+ extensional :: "'a set => ('a => 'b) set" where
"extensional A = {f. \<forall>x. x~:A --> f x = arbitrary}"
- "restrict" :: "['a => 'b, 'a set] => ('a => 'b)"
+definition
+ "restrict" :: "['a => 'b, 'a set] => ('a => 'b)" where
"restrict f A = (%x. if x \<in> A then f x else arbitrary)"
abbreviation
- funcset :: "['a set, 'b set] => ('a => 'b) set" (infixr "->" 60)
+ funcset :: "['a set, 'b set] => ('a => 'b) set"
+ (infixr "->" 60) where
"A -> B == Pi A (%_. B)"
notation (xsymbols)
@@ -43,7 +46,7 @@
"%x:A. f" == "CONST restrict (%x. f) A"
definition
- "compose" :: "['a set, 'b => 'c, 'a => 'b] => ('a => 'c)"
+ "compose" :: "['a set, 'b => 'c, 'a => 'b] => ('a => 'c)" where
"compose A g f = (\<lambda>x\<in>A. g (f x))"
@@ -142,7 +145,7 @@
the theorems belong here, or need at least @{term Hilbert_Choice}.*}
definition
- bij_betw :: "['a => 'b, 'a set, 'b set] => bool" -- {* bijective *}
+ bij_betw :: "['a => 'b, 'a set, 'b set] => bool" where -- {* bijective *}
"bij_betw f A B = (inj_on f A & f ` A = B)"
lemma bij_betw_imp_inj_on: "bij_betw f A B \<Longrightarrow> inj_on f A"
--- a/src/HOL/Library/GCD.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/GCD.thy Fri Nov 17 02:20:03 2006 +0100
@@ -22,7 +22,7 @@
"gcd (m, n) = (if n = 0 then m else gcd (n, m mod n))"
definition
- is_gcd :: "nat => nat => nat => bool" -- {* @{term gcd} as a relation *}
+ is_gcd :: "nat => nat => nat => bool" where -- {* @{term gcd} as a relation *}
"is_gcd p m n = (p dvd m \<and> p dvd n \<and>
(\<forall>d. d dvd m \<and> d dvd n --> d dvd p))"
--- a/src/HOL/Library/Infinite_Set.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Infinite_Set.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,7 +18,7 @@
*}
abbreviation
- infinite :: "'a set \<Rightarrow> bool"
+ infinite :: "'a set \<Rightarrow> bool" where
"infinite S == \<not> finite S"
text {*
@@ -349,17 +349,19 @@
*}
definition
- Inf_many :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "INF " 10)
+ Inf_many :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "INF " 10) where
"Inf_many P = infinite {x. P x}"
- Alm_all :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "MOST " 10)
+
+definition
+ Alm_all :: "('a \<Rightarrow> bool) \<Rightarrow> bool" (binder "MOST " 10) where
"Alm_all P = (\<not> (INF x. \<not> P x))"
notation (xsymbols)
- Inf_many (binder "\<exists>\<^sub>\<infinity>" 10)
+ Inf_many (binder "\<exists>\<^sub>\<infinity>" 10) and
Alm_all (binder "\<forall>\<^sub>\<infinity>" 10)
notation (HTML output)
- Inf_many (binder "\<exists>\<^sub>\<infinity>" 10)
+ Inf_many (binder "\<exists>\<^sub>\<infinity>" 10) and
Alm_all (binder "\<forall>\<^sub>\<infinity>" 10)
lemma INF_EX:
@@ -453,7 +455,7 @@
*}
definition
- atmost_one :: "'a set \<Rightarrow> bool"
+ atmost_one :: "'a set \<Rightarrow> bool" where
"atmost_one S = (\<forall>x y. x\<in>S \<and> y\<in>S \<longrightarrow> x=y)"
lemma atmost_one_empty: "S = {} \<Longrightarrow> atmost_one S"
--- a/src/HOL/Library/List_Prefix.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/List_Prefix.thy Fri Nov 17 02:20:03 2006 +0100
@@ -159,7 +159,7 @@
subsection {* Parallel lists *}
definition
- parallel :: "'a list => 'a list => bool" (infixl "\<parallel>" 50)
+ parallel :: "'a list => 'a list => bool" (infixl "\<parallel>" 50) where
"(xs \<parallel> ys) = (\<not> xs \<le> ys \<and> \<not> ys \<le> xs)"
lemma parallelI [intro]: "\<not> xs \<le> ys ==> \<not> ys \<le> xs ==> xs \<parallel> ys"
@@ -218,7 +218,7 @@
subsection {* Postfix order on lists *}
definition
- postfix :: "'a list => 'a list => bool" ("(_/ >>= _)" [51, 50] 50)
+ postfix :: "'a list => 'a list => bool" ("(_/ >>= _)" [51, 50] 50) where
"(xs >>= ys) = (\<exists>zs. xs = zs @ ys)"
lemma postfixI [intro?]: "xs = zs @ ys ==> xs >>= ys"
--- a/src/HOL/Library/Multiset.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Multiset.thy Fri Nov 17 02:20:03 2006 +0100
@@ -21,20 +21,23 @@
and [simp] = Rep_multiset_inject [symmetric]
definition
- Mempty :: "'a multiset" ("{#}")
+ Mempty :: "'a multiset" ("{#}") where
"{#} = Abs_multiset (\<lambda>a. 0)"
- single :: "'a => 'a multiset" ("{#_#}")
+definition
+ single :: "'a => 'a multiset" ("{#_#}") where
"{#a#} = Abs_multiset (\<lambda>b. if b = a then 1 else 0)"
- count :: "'a multiset => 'a => nat"
+definition
+ count :: "'a multiset => 'a => nat" where
"count = Rep_multiset"
- MCollect :: "'a multiset => ('a => bool) => 'a multiset"
+definition
+ MCollect :: "'a multiset => ('a => bool) => 'a multiset" where
"MCollect M P = Abs_multiset (\<lambda>x. if P x then Rep_multiset M x else 0)"
abbreviation
- Melem :: "'a => 'a multiset => bool" ("(_/ :# _)" [50, 51] 50)
+ Melem :: "'a => 'a multiset => bool" ("(_/ :# _)" [50, 51] 50) where
"a :# M == 0 < count M a"
syntax
@@ -43,7 +46,7 @@
"{#x:M. P#}" == "CONST MCollect M (\<lambda>x. P)"
definition
- set_of :: "'a multiset => 'a set"
+ set_of :: "'a multiset => 'a set" where
"set_of M = {x. x :# M}"
instance multiset :: (type) "{plus, minus, zero}" ..
@@ -55,7 +58,7 @@
size_def: "size M == setsum (count M) (set_of M)"
definition
- multiset_inter :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70)
+ multiset_inter :: "'a multiset \<Rightarrow> 'a multiset \<Rightarrow> 'a multiset" (infixl "#\<inter>" 70) where
"multiset_inter A B = A - (A - B)"
@@ -380,12 +383,13 @@
subsubsection {* Well-foundedness *}
definition
- mult1 :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set"
+ mult1 :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
"mult1 r =
{(N, M). \<exists>a M0 K. M = M0 + {#a#} \<and> N = M0 + K \<and>
(\<forall>b. b :# K --> (b, a) \<in> r)}"
- mult :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set"
+definition
+ mult :: "('a \<times> 'a) set => ('a multiset \<times> 'a multiset) set" where
"mult r = (mult1 r)\<^sup>+"
lemma not_less_empty [iff]: "(M, {#}) \<notin> mult1 r"
@@ -796,7 +800,7 @@
subsection {* Pointwise ordering induced by count *}
definition
- mset_le :: "['a multiset, 'a multiset] \<Rightarrow> bool" ("_ \<le># _" [50,51] 50)
+ mset_le :: "['a multiset, 'a multiset] \<Rightarrow> bool" ("_ \<le># _" [50,51] 50) where
"(xs \<le># ys) = (\<forall>a. count xs a \<le> count ys a)"
lemma mset_le_refl[simp]: "xs \<le># xs"
--- a/src/HOL/Library/NatPair.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/NatPair.thy Fri Nov 17 02:20:03 2006 +0100
@@ -16,7 +16,7 @@
*}
definition
- nat2_to_nat:: "(nat * nat) \<Rightarrow> nat"
+ nat2_to_nat:: "(nat * nat) \<Rightarrow> nat" where
"nat2_to_nat pair = (let (n,m) = pair in (n+m) * Suc (n+m) div 2 + n)"
lemma dvd2_a_x_suc_a: "2 dvd a * (Suc a)"
--- a/src/HOL/Library/Nat_Infinity.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Nat_Infinity.thy Fri Nov 17 02:20:03 2006 +0100
@@ -28,7 +28,7 @@
instance inat :: "{ord, zero}" ..
definition
- iSuc :: "inat => inat"
+ iSuc :: "inat => inat" where
"iSuc i = (case i of Fin n => Fin (Suc n) | \<infinity> => \<infinity>)"
defs (overloaded)
--- a/src/HOL/Library/OptionalSugar.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/OptionalSugar.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,7 +18,7 @@
(* aligning equations *)
notation (tab output)
- "op =" ("(_) \<^raw:}\putisatab\isa{\ >=\<^raw:}\putisatab\isa{> (_)" [50,49] 50)
+ "op =" ("(_) \<^raw:}\putisatab\isa{\ >=\<^raw:}\putisatab\isa{> (_)" [50,49] 50) and
"==" ("(_) \<^raw:}\putisatab\isa{\ >\<equiv>\<^raw:}\putisatab\isa{> (_)")
(* Let *)
--- a/src/HOL/Library/Parity.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Parity.thy Fri Nov 17 02:20:03 2006 +0100
@@ -22,7 +22,7 @@
even_nat_def: "even (x::nat) == even (int x)"
abbreviation
- odd :: "'a::even_odd => bool"
+ odd :: "'a::even_odd => bool" where
"odd x == \<not> even x"
--- a/src/HOL/Library/Permutation.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Permutation.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,7 +12,7 @@
perm :: "('a list * 'a list) set"
abbreviation
- perm_rel :: "'a list => 'a list => bool" ("_ <~~> _" [50, 50] 50)
+ perm_rel :: "'a list => 'a list => bool" ("_ <~~> _" [50, 50] 50) where
"x <~~> y == (x, y) \<in> perm"
inductive perm
--- a/src/HOL/Library/Primes.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Primes.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,10 +11,11 @@
begin
definition
- coprime :: "nat => nat => bool"
+ coprime :: "nat => nat => bool" where
"coprime m n = (gcd (m, n) = 1)"
- prime :: "nat \<Rightarrow> bool"
+definition
+ prime :: "nat \<Rightarrow> bool" where
"prime p = (1 < p \<and> (\<forall>m. m dvd p --> m = 1 \<or> m = p))"
--- a/src/HOL/Library/Quotient.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Quotient.thy Fri Nov 17 02:20:03 2006 +0100
@@ -75,7 +75,7 @@
*}
definition
- "class" :: "'a::equiv => 'a quot" ("\<lfloor>_\<rfloor>")
+ "class" :: "'a::equiv => 'a quot" ("\<lfloor>_\<rfloor>") where
"\<lfloor>a\<rfloor> = Abs_quot {x. a \<sim> x}"
theorem quot_exhaust: "\<exists>a. A = \<lfloor>a\<rfloor>"
@@ -134,7 +134,7 @@
subsection {* Picking representing elements *}
definition
- pick :: "'a::equiv quot => 'a"
+ pick :: "'a::equiv quot => 'a" where
"pick A = (SOME a. A = \<lfloor>a\<rfloor>)"
theorem pick_equiv [intro]: "pick \<lfloor>a\<rfloor> \<sim> a"
--- a/src/HOL/Library/SetsAndFunctions.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/SetsAndFunctions.thy Fri Nov 17 02:20:03 2006 +0100
@@ -52,14 +52,15 @@
set_one: "1::('a::one)set == {1}"
definition
- elt_set_plus :: "'a::plus => 'a set => 'a set" (infixl "+o" 70)
+ elt_set_plus :: "'a::plus => 'a set => 'a set" (infixl "+o" 70) where
"a +o B = {c. EX b:B. c = a + b}"
- elt_set_times :: "'a::times => 'a set => 'a set" (infixl "*o" 80)
+definition
+ elt_set_times :: "'a::times => 'a set => 'a set" (infixl "*o" 80) where
"a *o B = {c. EX b:B. c = a * b}"
abbreviation (input)
- elt_set_eq :: "'a => 'a set => bool" (infix "=o" 50)
+ elt_set_eq :: "'a => 'a set => bool" (infix "=o" 50) where
"x =o A == x : A"
instance "fun" :: (type,semigroup_add)semigroup_add
--- a/src/HOL/Library/State_Monad.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/State_Monad.thy Fri Nov 17 02:20:03 2006 +0100
@@ -68,12 +68,16 @@
definition
mbind :: "('a \<Rightarrow> 'b \<times> 'c) \<Rightarrow> ('b \<Rightarrow> 'c \<Rightarrow> 'd) \<Rightarrow> 'a \<Rightarrow> 'd"
- (infixl ">>=" 60)
+ (infixl ">>=" 60) where
"f >>= g = split g \<circ> f"
+
+definition
fcomp :: "('a \<Rightarrow> 'b) \<Rightarrow> ('b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'c"
- (infixl ">>" 60)
+ (infixl ">>" 60) where
"f >> g = g \<circ> f"
- run :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b"
+
+definition
+ run :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b" where
"run f = f"
print_ast_translation {*[
--- a/src/HOL/Library/While_Combinator.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/While_Combinator.thy Fri Nov 17 02:20:03 2006 +0100
@@ -36,7 +36,7 @@
done
definition
- while :: "('a => bool) => ('a => 'a) => 'a => 'a"
+ while :: "('a => bool) => ('a => 'a) => 'a => 'a" where
"while b c s = while_aux (b, c, s)"
lemma while_aux_unfold:
--- a/src/HOL/Library/Word.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Word.thy Fri Nov 17 02:20:03 2006 +0100
@@ -57,15 +57,15 @@
bitxor :: "bit => bit => bit" (infixr "bitxor" 30)
notation (xsymbols)
- bitnot ("\<not>\<^sub>b _" [40] 40)
- bitand (infixr "\<and>\<^sub>b" 35)
- bitor (infixr "\<or>\<^sub>b" 30)
+ bitnot ("\<not>\<^sub>b _" [40] 40) and
+ bitand (infixr "\<and>\<^sub>b" 35) and
+ bitor (infixr "\<or>\<^sub>b" 30) and
bitxor (infixr "\<oplus>\<^sub>b" 30)
notation (HTML output)
- bitnot ("\<not>\<^sub>b _" [40] 40)
- bitand (infixr "\<and>\<^sub>b" 35)
- bitor (infixr "\<or>\<^sub>b" 30)
+ bitnot ("\<not>\<^sub>b _" [40] 40) and
+ bitand (infixr "\<and>\<^sub>b" 35) and
+ bitor (infixr "\<or>\<^sub>b" 30) and
bitxor (infixr "\<oplus>\<^sub>b" 30)
primrec
@@ -142,11 +142,15 @@
qed
definition
- bv_msb :: "bit list => bit"
+ bv_msb :: "bit list => bit" where
"bv_msb w = (if w = [] then \<zero> else hd w)"
- bv_extend :: "[nat,bit,bit list]=>bit list"
+
+definition
+ bv_extend :: "[nat,bit,bit list]=>bit list" where
"bv_extend i b w = (replicate (i - length w) b) @ w"
- bv_not :: "bit list => bit list"
+
+definition
+ bv_not :: "bit list => bit list" where
"bv_not w = map bitnot w"
lemma bv_length_extend [simp]: "length w \<le> i ==> length (bv_extend i b w) = i"
@@ -177,7 +181,7 @@
by (induct w,simp_all)
definition
- bv_to_nat :: "bit list => nat"
+ bv_to_nat :: "bit list => nat" where
"bv_to_nat = foldl (%bn b. 2 * bn + bitval b) 0"
lemma bv_to_nat_Nil [simp]: "bv_to_nat [] = 0"
@@ -327,7 +331,7 @@
qed
definition
- norm_unsigned :: "bit list => bit list"
+ norm_unsigned :: "bit list => bit list" where
"norm_unsigned = rem_initial \<zero>"
lemma norm_unsigned_Nil [simp]: "norm_unsigned [] = []"
@@ -350,7 +354,7 @@
else nat_to_bv_helper (n div 2) ((if n mod 2 = 0 then \<zero> else \<one>)#bs)))"
definition
- nat_to_bv :: "nat => bit list"
+ nat_to_bv :: "nat => bit list" where
"nat_to_bv n = nat_to_bv_helper n []"
lemma nat_to_bv0 [simp]: "nat_to_bv 0 = []"
@@ -857,7 +861,7 @@
subsection {* Unsigned Arithmetic Operations *}
definition
- bv_add :: "[bit list, bit list ] => bit list"
+ bv_add :: "[bit list, bit list ] => bit list" where
"bv_add w1 w2 = nat_to_bv (bv_to_nat w1 + bv_to_nat w2)"
lemma bv_add_type1 [simp]: "bv_add (norm_unsigned w1) w2 = bv_add w1 w2"
@@ -909,7 +913,7 @@
qed
definition
- bv_mult :: "[bit list, bit list ] => bit list"
+ bv_mult :: "[bit list, bit list ] => bit list" where
"bv_mult w1 w2 = nat_to_bv (bv_to_nat w1 * bv_to_nat w2)"
lemma bv_mult_type1 [simp]: "bv_mult (norm_unsigned w1) w2 = bv_mult w1 w2"
@@ -969,7 +973,7 @@
lemmas [simp del] = norm_signed_Cons
definition
- int_to_bv :: "int => bit list"
+ int_to_bv :: "int => bit list" where
"int_to_bv n = (if 0 \<le> n
then norm_signed (\<zero>#nat_to_bv (nat n))
else norm_signed (bv_not (\<zero>#nat_to_bv (nat (-n- 1)))))"
@@ -1004,7 +1008,7 @@
qed
definition
- bv_to_int :: "bit list => int"
+ bv_to_int :: "bit list => int" where
"bv_to_int w =
(case bv_msb w of \<zero> => int (bv_to_nat w)
| \<one> => - int (bv_to_nat (bv_not w) + 1))"
@@ -1589,7 +1593,7 @@
subsubsection {* Conversion from unsigned to signed *}
definition
- utos :: "bit list => bit list"
+ utos :: "bit list => bit list" where
"utos w = norm_signed (\<zero> # w)"
lemma utos_type [simp]: "utos (norm_unsigned w) = utos w"
@@ -1613,7 +1617,7 @@
subsubsection {* Unary minus *}
definition
- bv_uminus :: "bit list => bit list"
+ bv_uminus :: "bit list => bit list" where
"bv_uminus w = int_to_bv (- bv_to_int w)"
lemma bv_uminus_type [simp]: "bv_uminus (norm_signed w) = bv_uminus w"
@@ -1712,7 +1716,7 @@
qed
definition
- bv_sadd :: "[bit list, bit list ] => bit list"
+ bv_sadd :: "[bit list, bit list ] => bit list" where
"bv_sadd w1 w2 = int_to_bv (bv_to_int w1 + bv_to_int w2)"
lemma bv_sadd_type1 [simp]: "bv_sadd (norm_signed w1) w2 = bv_sadd w1 w2"
@@ -1823,7 +1827,7 @@
qed
definition
- bv_sub :: "[bit list, bit list] => bit list"
+ bv_sub :: "[bit list, bit list] => bit list" where
"bv_sub w1 w2 = bv_sadd w1 (bv_uminus w2)"
lemma bv_sub_type1 [simp]: "bv_sub (norm_signed w1) w2 = bv_sub w1 w2"
@@ -1917,7 +1921,7 @@
qed
definition
- bv_smult :: "[bit list, bit list] => bit list"
+ bv_smult :: "[bit list, bit list] => bit list" where
"bv_smult w1 w2 = int_to_bv (bv_to_int w1 * bv_to_int w2)"
lemma bv_smult_type1 [simp]: "bv_smult (norm_signed w1) w2 = bv_smult w1 w2"
@@ -2203,11 +2207,15 @@
subsection {* Structural operations *}
definition
- bv_select :: "[bit list,nat] => bit"
+ bv_select :: "[bit list,nat] => bit" where
"bv_select w i = w ! (length w - 1 - i)"
- bv_chop :: "[bit list,nat] => bit list * bit list"
+
+definition
+ bv_chop :: "[bit list,nat] => bit list * bit list" where
"bv_chop w i = (let len = length w in (take (len - i) w,drop (len - i) w))"
- bv_slice :: "[bit list,nat*nat] => bit list"
+
+definition
+ bv_slice :: "[bit list,nat*nat] => bit list" where
"bv_slice w = (\<lambda>(b,e). fst (bv_chop (snd (bv_chop w (b+1))) e))"
lemma bv_select_rev:
@@ -2280,7 +2288,7 @@
by (auto simp add: bv_slice_def)
definition
- length_nat :: "nat => nat"
+ length_nat :: "nat => nat" where
"length_nat x = (LEAST n. x < 2 ^ n)"
lemma length_nat: "length (nat_to_bv n) = length_nat n"
@@ -2312,7 +2320,7 @@
done
definition
- length_int :: "int => nat"
+ length_int :: "int => nat" where
"length_int x =
(if 0 < x then Suc (length_nat (nat x))
else if x = 0 then 0
@@ -2546,7 +2554,7 @@
declare bv_to_nat_helper [simp del]
definition
- bv_mapzip :: "[bit => bit => bit,bit list, bit list] => bit list"
+ bv_mapzip :: "[bit => bit => bit,bit list, bit list] => bit list" where
"bv_mapzip f w1 w2 =
(let g = bv_extend (max (length w1) (length w2)) \<zero>
in map (split f) (zip (g w1) (g w2)))"
--- a/src/HOL/Library/Zorn.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Library/Zorn.thy Fri Nov 17 02:20:03 2006 +0100
@@ -16,16 +16,19 @@
*}
definition
- chain :: "'a set set => 'a set set set"
+ chain :: "'a set set => 'a set set set" where
"chain S = {F. F \<subseteq> S & (\<forall>x \<in> F. \<forall>y \<in> F. x \<subseteq> y | y \<subseteq> x)}"
- super :: "['a set set,'a set set] => 'a set set set"
+definition
+ super :: "['a set set,'a set set] => 'a set set set" where
"super S c = {d. d \<in> chain S & c \<subset> d}"
- maxchain :: "'a set set => 'a set set set"
+definition
+ maxchain :: "'a set set => 'a set set set" where
"maxchain S = {c. c \<in> chain S & super S c = {}}"
- succ :: "['a set set,'a set set] => 'a set set"
+definition
+ succ :: "['a set set,'a set set] => 'a set set" where
"succ S c =
(if c \<notin> chain S | c \<in> maxchain S
then c else SOME c'. c' \<in> super S c)"
--- a/src/HOL/List.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/List.thy Fri Nov 17 02:20:03 2006 +0100
@@ -43,7 +43,7 @@
splice :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
abbreviation
- upto:: "nat => nat => nat list" ("(1[_../_])")
+ upto:: "nat => nat => nat list" ("(1[_../_])") where
"[i..j] == [i..<(Suc j)]"
@@ -82,7 +82,7 @@
refer to the list version as @{text length}. *}
abbreviation
- length :: "'a list => nat"
+ length :: "'a list => nat" where
"length == size"
primrec
@@ -187,16 +187,21 @@
replicate_Suc: "replicate (Suc n) x = x # replicate n x"
definition
- rotate1 :: "'a list \<Rightarrow> 'a list"
- rotate1_def: "rotate1 xs = (case xs of [] \<Rightarrow> [] | x#xs \<Rightarrow> xs @ [x])"
- rotate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list"
- rotate_def: "rotate n = rotate1 ^ n"
- list_all2 :: "('a => 'b => bool) => 'a list => 'b list => bool"
- list_all2_def: "list_all2 P xs ys =
+ rotate1 :: "'a list \<Rightarrow> 'a list" where
+ "rotate1 xs = (case xs of [] \<Rightarrow> [] | x#xs \<Rightarrow> xs @ [x])"
+
+definition
+ rotate :: "nat \<Rightarrow> 'a list \<Rightarrow> 'a list" where
+ "rotate n = rotate1 ^ n"
+
+definition
+ list_all2 :: "('a => 'b => bool) => 'a list => 'b list => bool" where
+ "list_all2 P xs ys =
(length xs = length ys \<and> (\<forall>(x, y) \<in> set (zip xs ys). P x y))"
- sublist :: "'a list => nat set => 'a list"
- sublist_def: "sublist xs A =
- map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [0..<size xs]))"
+
+definition
+ sublist :: "'a list => nat set => 'a list" where
+ "sublist xs A = map fst (filter (\<lambda>p. snd p \<in> A) (zip xs [0..<size xs]))"
primrec
"splice [] ys = ys"
--- a/src/HOL/Map.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Map.thy Fri Nov 17 02:20:03 2006 +0100
@@ -19,34 +19,37 @@
"~=>" :: "[type, type] => type" (infixr "\<rightharpoonup>" 0)
abbreviation
- empty :: "'a ~=> 'b"
+ empty :: "'a ~=> 'b" where
"empty == %x. None"
definition
- map_comp :: "('b ~=> 'c) => ('a ~=> 'b) => ('a ~=> 'c)" (infixl "o'_m" 55)
+ map_comp :: "('b ~=> 'c) => ('a ~=> 'b) => ('a ~=> 'c)" (infixl "o'_m" 55) where
"f o_m g = (\<lambda>k. case g k of None \<Rightarrow> None | Some v \<Rightarrow> f v)"
notation (xsymbols)
map_comp (infixl "\<circ>\<^sub>m" 55)
definition
- map_add :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)" (infixl "++" 100)
+ map_add :: "('a ~=> 'b) => ('a ~=> 'b) => ('a ~=> 'b)" (infixl "++" 100) where
"m1 ++ m2 = (\<lambda>x. case m2 x of None => m1 x | Some y => Some y)"
- restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" (infixl "|`" 110)
+definition
+ restrict_map :: "('a ~=> 'b) => 'a set => ('a ~=> 'b)" (infixl "|`" 110) where
"m|`A = (\<lambda>x. if x : A then m x else None)"
notation (latex output)
restrict_map ("_\<restriction>\<^bsub>_\<^esub>" [111,110] 110)
definition
- dom :: "('a ~=> 'b) => 'a set"
+ dom :: "('a ~=> 'b) => 'a set" where
"dom m = {a. m a ~= None}"
- ran :: "('a ~=> 'b) => 'b set"
+definition
+ ran :: "('a ~=> 'b) => 'b set" where
"ran m = {b. EX a. m a = Some b}"
- map_le :: "('a ~=> 'b) => ('a ~=> 'b) => bool" (infix "\<subseteq>\<^sub>m" 50)
+definition
+ map_le :: "('a ~=> 'b) => ('a ~=> 'b) => bool" (infix "\<subseteq>\<^sub>m" 50) where
"(m\<^isub>1 \<subseteq>\<^sub>m m\<^isub>2) = (\<forall>a \<in> dom m\<^isub>1. m\<^isub>1 a = m\<^isub>2 a)"
consts
--- a/src/HOL/MicroJava/J/Example.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/MicroJava/J/Example.thy Fri Nov 17 02:20:03 2006 +0100
@@ -111,19 +111,21 @@
abbreviation
- NP :: xcpt
+ NP :: xcpt where
"NP == NullPointer"
- tprg ::"java_mb prog"
+abbreviation
+ tprg ::"java_mb prog" where
"tprg == [ObjectC, BaseC, ExtC, ClassCastC, NullPointerC, OutOfMemoryC]"
- obj1 :: obj
+abbreviation
+ obj1 :: obj where
"obj1 == (Ext, empty((vee, Base)\<mapsto>Bool False) ((vee, Ext )\<mapsto>Intg 0))"
- "s0 == Norm (empty, empty)"
- "s1 == Norm (empty(a\<mapsto>obj1),empty(e\<mapsto>Addr a))"
- "s2 == Norm (empty(a\<mapsto>obj1),empty(x\<mapsto>Null)(This\<mapsto>Addr a))"
- "s3 == (Some NP, empty(a\<mapsto>obj1),empty(e\<mapsto>Addr a))"
+abbreviation "s0 == Norm (empty, empty)"
+abbreviation "s1 == Norm (empty(a\<mapsto>obj1),empty(e\<mapsto>Addr a))"
+abbreviation "s2 == Norm (empty(a\<mapsto>obj1),empty(x\<mapsto>Null)(This\<mapsto>Addr a))"
+abbreviation "s3 == (Some NP, empty(a\<mapsto>obj1),empty(e\<mapsto>Addr a))"
ML {* bind_thm ("map_of_Cons", hd (tl (thms "map_of.simps"))) *}
lemma map_of_Cons1 [simp]: "map_of ((aa,bb)#ps) aa = Some bb"
--- a/src/HOL/MicroJava/JVM/JVMListExample.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/MicroJava/JVM/JVMListExample.thy Fri Nov 17 02:20:03 2006 +0100
@@ -124,14 +124,14 @@
and instr :: eq ..
definition
- arbitrary_val :: val
+ arbitrary_val :: val where
[symmetric, code inline]: "arbitrary_val = arbitrary"
- arbitrary_cname :: cname
+definition
+ arbitrary_cname :: cname where
[symmetric, code inline]: "arbitrary_cname = arbitrary"
-definition
- "unit' = Unit"
- "object' = Object"
+definition "unit' = Unit"
+definition "object' = Object"
code_axioms
arbitrary_val \<equiv> unit'
@@ -153,7 +153,7 @@
"test_loc p v l = (if p l then v l else test_loc p v (incr l))"
definition
- new_addr' :: "(loc \<Rightarrow> (cname \<times> (vname \<times> cname \<Rightarrow> val option)) option) \<Rightarrow> loc \<times> val option"
+ new_addr' :: "(loc \<Rightarrow> (cname \<times> (vname \<times> cname \<Rightarrow> val option)) option) \<Rightarrow> loc \<times> val option" where
"new_addr' hp =
test_loc (\<lambda>i. hp (Loc i) = None) (\<lambda>i. (Loc i, None)) zero_loc"
--- a/src/HOL/Nominal/Examples/Lam_Funs.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Nominal/Examples/Lam_Funs.thy Fri Nov 17 02:20:03 2006 +0100
@@ -1,4 +1,4 @@
-(* $Id: *)
+(* $Id$ *)
theory Lam_Funs
imports Nominal
@@ -73,7 +73,7 @@
"subst_Lam b t \<equiv> \<lambda>a _ r. Lam [a].r"
abbreviation
- subst_syn :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 900)
+ subst_syn :: "lam \<Rightarrow> name \<Rightarrow> lam \<Rightarrow> lam" ("_[_::=_]" [100,100,100] 900) where
"t'[b::=t] \<equiv> (lam_rec (subst_Var b t) (subst_App b t) (subst_Lam b t)) t'"
(* FIXME: this lemma needs to be in Nominal.thy *)
@@ -200,7 +200,7 @@
"psubst_Lam \<theta> \<equiv> \<lambda>a _ r. Lam [a].r"
abbreviation
- psubst_lam :: "lam \<Rightarrow> (name\<times>lam) list \<Rightarrow> lam" ("_[<_>]" [100,100] 900)
+ psubst_lam :: "lam \<Rightarrow> (name\<times>lam) list \<Rightarrow> lam" ("_[<_>]" [100,100] 900) where
"t[<\<theta>>] \<equiv> (lam_rec (psubst_Var \<theta>) (psubst_App \<theta>) (psubst_Lam \<theta>)) t"
lemma fin_supp_psubst:
@@ -241,4 +241,4 @@
apply(simp add: psubst_Lam_def)
done
-end
\ No newline at end of file
+end
--- a/src/HOL/Nominal/Examples/SN.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Nominal/Examples/SN.thy Fri Nov 17 02:20:03 2006 +0100
@@ -344,7 +344,7 @@
qed
abbreviation
- "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [80,80] 80)
+ "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [80,80] 80) where
"\<Gamma>1 \<lless> \<Gamma>2 \<equiv> \<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow> (a,\<sigma>)\<in>set \<Gamma>2"
lemma weakening:
--- a/src/HOL/Nominal/Examples/Weakening.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Nominal/Examples/Weakening.thy Fri Nov 17 02:20:03 2006 +0100
@@ -193,7 +193,7 @@
text {* definition of a subcontext *}
abbreviation
- "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [80,80] 80)
+ "sub" :: "(name\<times>ty) list \<Rightarrow> (name\<times>ty) list \<Rightarrow> bool" (" _ \<lless> _ " [80,80] 80) where
"\<Gamma>1 \<lless> \<Gamma>2 \<equiv> \<forall>a \<sigma>. (a,\<sigma>)\<in>set \<Gamma>1 \<longrightarrow> (a,\<sigma>)\<in>set \<Gamma>2"
text {* Now it comes: The Weakening Lemma *}
--- a/src/HOL/NumberTheory/BijectionRel.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/BijectionRel.thy Fri Nov 17 02:20:03 2006 +0100
@@ -30,13 +30,15 @@
*}
definition
- bijP :: "('a => 'a => bool) => 'a set => bool"
+ bijP :: "('a => 'a => bool) => 'a set => bool" where
"bijP P F = (\<forall>a b. a \<in> F \<and> P a b --> b \<in> F)"
- uniqP :: "('a => 'a => bool) => bool"
+definition
+ uniqP :: "('a => 'a => bool) => bool" where
"uniqP P = (\<forall>a b c d. P a b \<and> P c d --> (a = c) = (b = d))"
- symP :: "('a => 'a => bool) => bool"
+definition
+ symP :: "('a => 'a => bool) => bool" where
"symP P = (\<forall>a b. P a b = P b a)"
consts
--- a/src/HOL/NumberTheory/Chinese.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/Chinese.thy Fri Nov 17 02:20:03 2006 +0100
@@ -32,32 +32,37 @@
"funsum f i (Suc n) = f (Suc (i + n)) + funsum f i n"
definition
- m_cond :: "nat => (nat => int) => bool"
+ m_cond :: "nat => (nat => int) => bool" where
"m_cond n mf =
((\<forall>i. i \<le> n --> 0 < mf i) \<and>
(\<forall>i j. i \<le> n \<and> j \<le> n \<and> i \<noteq> j --> zgcd (mf i, mf j) = 1))"
- km_cond :: "nat => (nat => int) => (nat => int) => bool"
+definition
+ km_cond :: "nat => (nat => int) => (nat => int) => bool" where
"km_cond n kf mf = (\<forall>i. i \<le> n --> zgcd (kf i, mf i) = 1)"
+definition
lincong_sol ::
- "nat => (nat => int) => (nat => int) => (nat => int) => int => bool"
+ "nat => (nat => int) => (nat => int) => (nat => int) => int => bool" where
"lincong_sol n kf bf mf x = (\<forall>i. i \<le> n --> zcong (kf i * x) (bf i) (mf i))"
- mhf :: "(nat => int) => nat => nat => int"
+definition
+ mhf :: "(nat => int) => nat => nat => int" where
"mhf mf n i =
(if i = 0 then funprod mf (Suc 0) (n - Suc 0)
else if i = n then funprod mf 0 (n - Suc 0)
else funprod mf 0 (i - Suc 0) * funprod mf (Suc i) (n - Suc 0 - i))"
+definition
xilin_sol ::
- "nat => nat => (nat => int) => (nat => int) => (nat => int) => int"
+ "nat => nat => (nat => int) => (nat => int) => (nat => int) => int" where
"xilin_sol i n kf bf mf =
(if 0 < n \<and> i \<le> n \<and> m_cond n mf \<and> km_cond n kf mf then
(SOME x. 0 \<le> x \<and> x < mf i \<and> zcong (kf i * mhf mf n i * x) (bf i) (mf i))
else 0)"
- x_sol :: "nat => (nat => int) => (nat => int) => (nat => int) => int"
+definition
+ x_sol :: "nat => (nat => int) => (nat => int) => (nat => int) => int" where
"x_sol n kf bf mf = funsum (\<lambda>i. xilin_sol i n kf bf mf * mhf mf n i) 0 n"
--- a/src/HOL/NumberTheory/Euler.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/Euler.thy Fri Nov 17 02:20:03 2006 +0100
@@ -8,10 +8,11 @@
theory Euler imports Residues EvenOdd begin
definition
- MultInvPair :: "int => int => int => int set"
+ MultInvPair :: "int => int => int => int set" where
"MultInvPair a p j = {StandardRes p j, StandardRes p (a * (MultInv p j))}"
- SetS :: "int => int => int set set"
+definition
+ SetS :: "int => int => int set set" where
"SetS a p = (MultInvPair a p ` SRStar p)"
--- a/src/HOL/NumberTheory/EulerFermat.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/EulerFermat.thy Fri Nov 17 02:20:03 2006 +0100
@@ -38,25 +38,30 @@
else {})"
definition
- norRRset :: "int => int set"
+ norRRset :: "int => int set" where
"norRRset m = BnorRset (m - 1, m)"
- noXRRset :: "int => int => int set"
+definition
+ noXRRset :: "int => int => int set" where
"noXRRset m x = (\<lambda>a. a * x) ` norRRset m"
- phi :: "int => nat"
+definition
+ phi :: "int => nat" where
"phi m = card (norRRset m)"
- is_RRset :: "int set => int => bool"
+definition
+ is_RRset :: "int set => int => bool" where
"is_RRset A m = (A \<in> RsetR m \<and> card A = phi m)"
- RRset2norRR :: "int set => int => int => int"
+definition
+ RRset2norRR :: "int set => int => int => int" where
"RRset2norRR A m a =
(if 1 < m \<and> is_RRset A m \<and> a \<in> A then
SOME b. zcong a b m \<and> b \<in> norRRset m
else 0)"
- zcongm :: "int => int => int => bool"
+definition
+ zcongm :: "int => int => int => bool" where
"zcongm m = (\<lambda>a b. zcong a b m)"
lemma abs_eq_1_iff [iff]: "(abs z = (1::int)) = (z = 1 \<or> z = -1)"
--- a/src/HOL/NumberTheory/EvenOdd.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/EvenOdd.thy Fri Nov 17 02:20:03 2006 +0100
@@ -8,9 +8,11 @@
theory EvenOdd imports Int2 begin
definition
- zOdd :: "int set"
+ zOdd :: "int set" where
"zOdd = {x. \<exists>k. x = 2 * k + 1}"
- zEven :: "int set"
+
+definition
+ zEven :: "int set" where
"zEven = {x. \<exists>k. x = 2 * k}"
subsection {* Some useful properties about even and odd *}
--- a/src/HOL/NumberTheory/Factorization.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/Factorization.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,7 +12,7 @@
subsection {* Definitions *}
definition
- primel :: "nat list => bool "
+ primel :: "nat list => bool" where
"primel xs = (\<forall>p \<in> set xs. prime p)"
consts
--- a/src/HOL/NumberTheory/Gauss.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/Gauss.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,22 +18,27 @@
begin
definition
- A :: "int set"
+ A :: "int set" where
"A = {(x::int). 0 < x & x \<le> ((p - 1) div 2)}"
- B :: "int set"
+definition
+ B :: "int set" where
"B = (%x. x * a) ` A"
- C :: "int set"
+definition
+ C :: "int set" where
"C = StandardRes p ` B"
- D :: "int set"
+definition
+ D :: "int set" where
"D = C \<inter> {x. x \<le> ((p - 1) div 2)}"
- E :: "int set"
+definition
+ E :: "int set" where
"E = C \<inter> {x. ((p - 1) div 2) < x}"
- F :: "int set"
+definition
+ F :: "int set" where
"F = (%x. (p - x)) ` E"
--- a/src/HOL/NumberTheory/Int2.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/Int2.thy Fri Nov 17 02:20:03 2006 +0100
@@ -8,7 +8,7 @@
theory Int2 imports Finite2 WilsonRuss begin
definition
- MultInv :: "int => int => int"
+ MultInv :: "int => int => int" where
"MultInv p x = x ^ nat (p - 2)"
--- a/src/HOL/NumberTheory/IntPrimes.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/IntPrimes.thy Fri Nov 17 02:20:03 2006 +0100
@@ -32,16 +32,19 @@
t, t' - (r' div r) * t))"
definition
- zgcd :: "int * int => int"
+ zgcd :: "int * int => int" where
"zgcd = (\<lambda>(x,y). int (gcd (nat (abs x), nat (abs y))))"
- zprime :: "int \<Rightarrow> bool"
+definition
+ zprime :: "int \<Rightarrow> bool" where
"zprime p = (1 < p \<and> (\<forall>m. 0 <= m & m dvd p --> m = 1 \<or> m = p))"
- xzgcd :: "int => int => int * int * int"
+definition
+ xzgcd :: "int => int => int * int * int" where
"xzgcd m n = xzgcda (m, n, m, n, 1, 0, 0, 1)"
- zcong :: "int => int => int => bool" ("(1[_ = _] '(mod _'))")
+definition
+ zcong :: "int => int => int => bool" ("(1[_ = _] '(mod _'))") where
"[a = b] (mod m) = (m dvd (a - b))"
--- a/src/HOL/NumberTheory/Quadratic_Reciprocity.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/Quadratic_Reciprocity.thy Fri Nov 17 02:20:03 2006 +0100
@@ -168,25 +168,31 @@
begin
definition
- P_set :: "int set"
+ P_set :: "int set" where
"P_set = {x. 0 < x & x \<le> ((p - 1) div 2) }"
- Q_set :: "int set"
+definition
+ Q_set :: "int set" where
"Q_set = {x. 0 < x & x \<le> ((q - 1) div 2) }"
- S :: "(int * int) set"
+definition
+ S :: "(int * int) set" where
"S = P_set <*> Q_set"
- S1 :: "(int * int) set"
+definition
+ S1 :: "(int * int) set" where
"S1 = { (x, y). (x, y):S & ((p * y) < (q * x)) }"
- S2 :: "(int * int) set"
+definition
+ S2 :: "(int * int) set" where
"S2 = { (x, y). (x, y):S & ((q * x) < (p * y)) }"
- f1 :: "int => (int * int) set"
+definition
+ f1 :: "int => (int * int) set" where
"f1 j = { (j1, y). (j1, y):S & j1 = j & (y \<le> (q * j) div p) }"
- f2 :: "int => (int * int) set"
+definition
+ f2 :: "int => (int * int) set" where
"f2 j = { (x, j1). (x, j1):S & j1 = j & (x \<le> (p * j) div q) }"
lemma p_fact: "0 < (p - 1) div 2"
--- a/src/HOL/NumberTheory/Residues.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/Residues.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,24 +12,29 @@
quadratic residues, and prove some basic properties. *}
definition
- ResSet :: "int => int set => bool"
+ ResSet :: "int => int set => bool" where
"ResSet m X = (\<forall>y1 y2. (y1 \<in> X & y2 \<in> X & [y1 = y2] (mod m) --> y1 = y2))"
- StandardRes :: "int => int => int"
+definition
+ StandardRes :: "int => int => int" where
"StandardRes m x = x mod m"
- QuadRes :: "int => int => bool"
+definition
+ QuadRes :: "int => int => bool" where
"QuadRes m x = (\<exists>y. ([(y ^ 2) = x] (mod m)))"
- Legendre :: "int => int => int"
+definition
+ Legendre :: "int => int => int" where
"Legendre a p = (if ([a = 0] (mod p)) then 0
else if (QuadRes p a) then 1
else -1)"
- SR :: "int => int set"
+definition
+ SR :: "int => int set" where
"SR p = {x. (0 \<le> x) & (x < p)}"
- SRStar :: "int => int set"
+definition
+ SRStar :: "int => int set" where
"SRStar p = {x. (0 < x) & (x < p)}"
--- a/src/HOL/NumberTheory/WilsonBij.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/WilsonBij.thy Fri Nov 17 02:20:03 2006 +0100
@@ -18,9 +18,11 @@
subsection {* Definitions and lemmas *}
definition
- reciR :: "int => int => int => bool"
+ reciR :: "int => int => int => bool" where
"reciR p = (\<lambda>a b. zcong (a * b) 1 p \<and> 1 < a \<and> a < p - 1 \<and> 1 < b \<and> b < p - 1)"
- inv :: "int => int => int"
+
+definition
+ inv :: "int => int => int" where
"inv p a =
(if zprime p \<and> 0 < a \<and> a < p then
(SOME x. 0 \<le> x \<and> x < p \<and> zcong (a * x) 1 p)
--- a/src/HOL/NumberTheory/WilsonRuss.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/NumberTheory/WilsonRuss.thy Fri Nov 17 02:20:03 2006 +0100
@@ -16,7 +16,7 @@
subsection {* Definitions and lemmas *}
definition
- inv :: "int => int => int"
+ inv :: "int => int => int" where
"inv p a = (a^(nat (p - 2))) mod p"
consts
--- a/src/HOL/Orderings.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Orderings.thy Fri Nov 17 02:20:03 2006 +0100
@@ -19,23 +19,25 @@
begin
notation
- less_eq ("op \<^loc><=")
- less_eq ("(_/ \<^loc><= _)" [51, 51] 50)
- less ("op \<^loc><")
+ less_eq ("op \<^loc><=") and
+ less_eq ("(_/ \<^loc><= _)" [51, 51] 50) and
+ less ("op \<^loc><") and
less ("(_/ \<^loc>< _)" [51, 51] 50)
notation (xsymbols)
- less_eq ("op \<^loc>\<le>")
+ less_eq ("op \<^loc>\<le>") and
less_eq ("(_/ \<^loc>\<le> _)" [51, 51] 50)
notation (HTML output)
- less_eq ("op \<^loc>\<le>")
+ less_eq ("op \<^loc>\<le>") and
less_eq ("(_/ \<^loc>\<le> _)" [51, 51] 50)
abbreviation (input)
- greater (infix "\<^loc>>" 50)
+ greater (infix "\<^loc>>" 50) where
"x \<^loc>> y \<equiv> y \<^loc>< x"
- greater_eq (infix "\<^loc>>=" 50)
+
+abbreviation (input)
+ greater_eq (infix "\<^loc>>=" 50) where
"x \<^loc>>= y \<equiv> y \<^loc><= x"
notation (xsymbols)
@@ -44,23 +46,25 @@
end
notation
- less_eq ("op <=")
- less_eq ("(_/ <= _)" [51, 51] 50)
- less ("op <")
+ less_eq ("op <=") and
+ less_eq ("(_/ <= _)" [51, 51] 50) and
+ less ("op <") and
less ("(_/ < _)" [51, 51] 50)
notation (xsymbols)
- less_eq ("op \<le>")
+ less_eq ("op \<le>") and
less_eq ("(_/ \<le> _)" [51, 51] 50)
notation (HTML output)
- less_eq ("op \<le>")
+ less_eq ("op \<le>") and
less_eq ("(_/ \<le> _)" [51, 51] 50)
abbreviation (input)
- greater (infix ">" 50)
+ greater (infix ">" 50) where
"x > y \<equiv> y < x"
- greater_eq (infix ">=" 50)
+
+abbreviation (input)
+ greater_eq (infix ">=" 50) where
"x >= y \<equiv> y <= x"
notation (xsymbols)
@@ -78,11 +82,11 @@
begin
abbreviation (input)
- greater_eq (infixl "\<sqsupseteq>" 50)
+ greater_eq (infixl "\<sqsupseteq>" 50) where
"x \<sqsupseteq> y \<equiv> y \<sqsubseteq> x"
abbreviation (input)
- greater (infixl "\<sqsupset>" 50)
+ greater (infixl "\<sqsupset>" 50) where
"x \<sqsupset> y \<equiv> y \<sqsubset> x"
text {* Reflexivity. *}
@@ -202,8 +206,6 @@
locale linorder = partial_order +
assumes linear: "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
-
-context linorder
begin
lemma trichotomy: "x \<sqsubset> y \<or> x = y \<or> y \<sqsubset> x"
@@ -259,9 +261,11 @@
(* min/max *)
definition
- min :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ min :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"min a b = (if a \<sqsubseteq> b then a else b)"
- max :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+
+definition
+ max :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"max a b = (if a \<sqsubseteq> b then b else a)"
lemma min_le_iff_disj:
--- a/src/HOL/Product_Type.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Product_Type.thy Fri Nov 17 02:20:03 2006 +0100
@@ -110,7 +110,8 @@
Sigma_def: "Sigma A B == UN x:A. UN y:B x. {Pair x y}"
abbreviation
- Times :: "['a set, 'b set] => ('a * 'b) set" (infixr "<*>" 80)
+ Times :: "['a set, 'b set] => ('a * 'b) set"
+ (infixr "<*>" 80) where
"A <*> B == Sigma A (%_. B)"
notation (xsymbols)
--- a/src/HOL/Real/ContNotDenum.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/ContNotDenum.thy Fri Nov 17 02:20:03 2006 +0100
@@ -37,7 +37,7 @@
subsubsection {* Definition *}
definition
- closed_int :: "real \<Rightarrow> real \<Rightarrow> real set"
+ closed_int :: "real \<Rightarrow> real \<Rightarrow> real set" where
"closed_int x y = {z. x \<le> z \<and> z \<le> y}"
subsubsection {* Properties *}
--- a/src/HOL/Real/Float.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/Float.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,9 +11,11 @@
begin
definition
- pow2 :: "int \<Rightarrow> real"
+ pow2 :: "int \<Rightarrow> real" where
"pow2 a = (if (0 <= a) then (2^(nat a)) else (inverse (2^(nat (-a)))))"
- float :: "int * int \<Rightarrow> real"
+
+definition
+ float :: "int * int \<Rightarrow> real" where
"float x = real (fst x) * pow2 (snd x)"
lemma pow2_0[simp]: "pow2 0 = 1"
@@ -99,9 +101,11 @@
by (simp add: float_def ring_eq_simps)
definition
- int_of_real :: "real \<Rightarrow> int"
+ int_of_real :: "real \<Rightarrow> int" where
"int_of_real x = (SOME y. real y = x)"
- real_is_int :: "real \<Rightarrow> bool"
+
+definition
+ real_is_int :: "real \<Rightarrow> bool" where
"real_is_int x = (EX (u::int). x = real u)"
lemma real_is_int_def2: "real_is_int x = (x = real (int_of_real x))"
--- a/src/HOL/Real/HahnBanach/Bounds.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/HahnBanach/Bounds.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,7 +15,7 @@
lemmas [elim?] = lub.least lub.upper
definition
- the_lub :: "'a::order set \<Rightarrow> 'a"
+ the_lub :: "'a::order set \<Rightarrow> 'a" where
"the_lub A = The (lub A)"
notation (xsymbols)
--- a/src/HOL/Real/HahnBanach/FunctionOrder.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/HahnBanach/FunctionOrder.thy Fri Nov 17 02:20:03 2006 +0100
@@ -23,7 +23,7 @@
types 'a graph = "('a \<times> real) set"
definition
- graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph"
+ graph :: "'a set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a graph" where
"graph F f = {(x, f x) | x. x \<in> F}"
lemma graphI [intro]: "x \<in> F \<Longrightarrow> (x, f x) \<in> graph F f"
@@ -66,10 +66,11 @@
*}
definition
- "domain" :: "'a graph \<Rightarrow> 'a set"
+ "domain" :: "'a graph \<Rightarrow> 'a set" where
"domain g = {x. \<exists>y. (x, y) \<in> g}"
- funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)"
+definition
+ funct :: "'a graph \<Rightarrow> ('a \<Rightarrow> real)" where
"funct g = (\<lambda>x. (SOME y. (x, y) \<in> g))"
text {*
@@ -104,7 +105,7 @@
definition
norm_pres_extensions ::
"'a::{plus, minus, zero} set \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> real)
- \<Rightarrow> 'a graph set"
+ \<Rightarrow> 'a graph set" where
"norm_pres_extensions E p F f
= {g. \<exists>H h. g = graph H h
\<and> linearform H h
--- a/src/HOL/Real/HahnBanach/Subspace.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/HahnBanach/Subspace.thy Fri Nov 17 02:20:03 2006 +0100
@@ -131,7 +131,7 @@
*}
definition
- lin :: "('a::{minus, plus, zero}) \<Rightarrow> 'a set"
+ lin :: "('a::{minus, plus, zero}) \<Rightarrow> 'a set" where
"lin x = {a \<cdot> x | a. True}"
lemma linI [intro]: "y = a \<cdot> x \<Longrightarrow> y \<in> lin x"
--- a/src/HOL/Real/Lubs.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/Lubs.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,23 +13,27 @@
text{*Thanks to suggestions by James Margetson*}
definition
-
- setle :: "['a set, 'a::ord] => bool" (infixl "*<=" 70)
+ setle :: "['a set, 'a::ord] => bool" (infixl "*<=" 70) where
"S *<= x = (ALL y: S. y <= x)"
- setge :: "['a::ord, 'a set] => bool" (infixl "<=*" 70)
+definition
+ setge :: "['a::ord, 'a set] => bool" (infixl "<=*" 70) where
"x <=* S = (ALL y: S. x <= y)"
- leastP :: "['a =>bool,'a::ord] => bool"
+definition
+ leastP :: "['a =>bool,'a::ord] => bool" where
"leastP P x = (P x & x <=* Collect P)"
- isUb :: "['a set, 'a set, 'a::ord] => bool"
+definition
+ isUb :: "['a set, 'a set, 'a::ord] => bool" where
"isUb R S x = (S *<= x & x: R)"
- isLub :: "['a set, 'a set, 'a::ord] => bool"
+definition
+ isLub :: "['a set, 'a set, 'a::ord] => bool" where
"isLub R S x = leastP (isUb R S) x"
- ubs :: "['a set, 'a::ord set] => 'a set"
+definition
+ ubs :: "['a set, 'a::ord set] => 'a set" where
"ubs R S = Collect (isUb R S)"
--- a/src/HOL/Real/PReal.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/PReal.thy Fri Nov 17 02:20:03 2006 +0100
@@ -28,7 +28,7 @@
definition
- cut :: "rat set => bool"
+ cut :: "rat set => bool" where
"cut A = ({} \<subset> A &
A < {r. 0 < r} &
(\<forall>y \<in> A. ((\<forall>z. 0<z & z < y --> z \<in> A) & (\<exists>u \<in> A. y < u))))"
@@ -56,22 +56,27 @@
instance preal :: "{ord, plus, minus, times, inverse}" ..
definition
- preal_of_rat :: "rat => preal"
+ preal_of_rat :: "rat => preal" where
"preal_of_rat q = Abs_preal {x::rat. 0 < x & x < q}"
- psup :: "preal set => preal"
+definition
+ psup :: "preal set => preal" where
"psup P = Abs_preal (\<Union>X \<in> P. Rep_preal X)"
- add_set :: "[rat set,rat set] => rat set"
+definition
+ add_set :: "[rat set,rat set] => rat set" where
"add_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x + y}"
- diff_set :: "[rat set,rat set] => rat set"
+definition
+ diff_set :: "[rat set,rat set] => rat set" where
"diff_set A B = {w. \<exists>x. 0 < w & 0 < x & x \<notin> B & x + w \<in> A}"
- mult_set :: "[rat set,rat set] => rat set"
+definition
+ mult_set :: "[rat set,rat set] => rat set" where
"mult_set A B = {w. \<exists>x \<in> A. \<exists>y \<in> B. w = x * y}"
- inverse_set :: "rat set => rat set"
+definition
+ inverse_set :: "rat set => rat set" where
"inverse_set A = {x. \<exists>y. 0 < x & x < y & inverse y \<notin> A}"
--- a/src/HOL/Real/RComplete.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/RComplete.thy Fri Nov 17 02:20:03 2006 +0100
@@ -432,18 +432,19 @@
subsection{*Floor and Ceiling Functions from the Reals to the Integers*}
definition
- floor :: "real => int"
+ floor :: "real => int" where
"floor r = (LEAST n::int. r < real (n+1))"
- ceiling :: "real => int"
+definition
+ ceiling :: "real => int" where
"ceiling r = - floor (- r)"
notation (xsymbols)
- floor ("\<lfloor>_\<rfloor>")
+ floor ("\<lfloor>_\<rfloor>") and
ceiling ("\<lceil>_\<rceil>")
notation (HTML output)
- floor ("\<lfloor>_\<rfloor>")
+ floor ("\<lfloor>_\<rfloor>") and
ceiling ("\<lceil>_\<rceil>")
@@ -933,9 +934,11 @@
subsection {* Versions for the natural numbers *}
definition
- natfloor :: "real => nat"
+ natfloor :: "real => nat" where
"natfloor x = nat(floor x)"
- natceiling :: "real => nat"
+
+definition
+ natceiling :: "real => nat" where
"natceiling x = nat(ceiling x)"
lemma natfloor_zero [simp]: "natfloor 0 = 0"
--- a/src/HOL/Real/Rational.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/Rational.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,10 +15,11 @@
subsubsection {* Equivalence of fractions *}
definition
- fraction :: "(int \<times> int) set"
+ fraction :: "(int \<times> int) set" where
"fraction = {x. snd x \<noteq> 0}"
- ratrel :: "((int \<times> int) \<times> (int \<times> int)) set"
+definition
+ ratrel :: "((int \<times> int) \<times> (int \<times> int)) set" where
"ratrel = {(x,y). snd x \<noteq> 0 \<and> snd y \<noteq> 0 \<and> fst x * snd y = fst y * snd x}"
lemma fraction_iff [simp]: "(x \<in> fraction) = (snd x \<noteq> 0)"
@@ -79,12 +80,12 @@
definition
- Fract :: "int \<Rightarrow> int \<Rightarrow> rat"
+ Fract :: "int \<Rightarrow> int \<Rightarrow> rat" where
"Fract a b = Abs_Rat (ratrel``{(a,b)})"
theorem Rat_cases [case_names Fract, cases type: rat]:
- "(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C"
-by (cases q, clarsimp simp add: Fract_def Rat_def fraction_def quotient_def)
+ "(!!a b. q = Fract a b ==> b \<noteq> 0 ==> C) ==> C"
+ by (cases q) (clarsimp simp add: Fract_def Rat_def fraction_def quotient_def)
theorem Rat_induct [case_names Fract, induct type: rat]:
"(!!a b. b \<noteq> 0 ==> P (Fract a b)) ==> P q"
--- a/src/HOL/Real/RealDef.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/RealDef.thy Fri Nov 17 02:20:03 2006 +0100
@@ -14,7 +14,7 @@
begin
definition
- realrel :: "((preal * preal) * (preal * preal)) set"
+ realrel :: "((preal * preal) * (preal * preal)) set" where
"realrel = {p. \<exists>x1 y1 x2 y2. p = ((x1,y1),(x2,y2)) & x1+y2 = x2+y1}"
typedef (Real) real = "UNIV//realrel"
@@ -26,7 +26,7 @@
(** these don't use the overloaded "real" function: users don't see them **)
- real_of_preal :: "preal => real"
+ real_of_preal :: "preal => real" where
"real_of_preal m = Abs_Real(realrel``{(m + preal_of_rat 1, preal_of_rat 1)})"
consts
--- a/src/HOL/Real/RealVector.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Real/RealVector.thy Fri Nov 17 02:20:03 2006 +0100
@@ -41,11 +41,11 @@
scaleR :: "real \<Rightarrow> 'a \<Rightarrow> 'a::scaleR" (infixr "*#" 75)
abbreviation
- divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a::scaleR" (infixl "'/#" 70)
+ divideR :: "'a \<Rightarrow> real \<Rightarrow> 'a::scaleR" (infixl "'/#" 70) where
"x /# r == inverse r *# x"
notation (xsymbols)
- scaleR (infixr "*\<^sub>R" 75)
+ scaleR (infixr "*\<^sub>R" 75) and
divideR (infixl "'/\<^sub>R" 70)
instance real :: scaleR ..
@@ -175,7 +175,7 @@
@{term of_real} *}
definition
- of_real :: "real \<Rightarrow> 'a::real_algebra_1"
+ of_real :: "real \<Rightarrow> 'a::real_algebra_1" where
"of_real r = r *# 1"
lemma scaleR_conv_of_real: "r *# x = of_real r * x"
@@ -250,7 +250,7 @@
subsection {* The Set of Real Numbers *}
definition
- Reals :: "'a::real_algebra_1 set"
+ Reals :: "'a::real_algebra_1 set" where
"Reals \<equiv> range of_real"
notation (xsymbols)
--- a/src/HOL/Relation.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Relation.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,54 +13,69 @@
subsection {* Definitions *}
definition
- converse :: "('a * 'b) set => ('b * 'a) set" ("(_^-1)" [1000] 999)
+ converse :: "('a * 'b) set => ('b * 'a) set"
+ ("(_^-1)" [1000] 999) where
"r^-1 == {(y, x). (x, y) : r}"
notation (xsymbols)
converse ("(_\<inverse>)" [1000] 999)
definition
- rel_comp :: "[('b * 'c) set, ('a * 'b) set] => ('a * 'c) set" (infixr "O" 75)
+ rel_comp :: "[('b * 'c) set, ('a * 'b) set] => ('a * 'c) set"
+ (infixr "O" 75) where
"r O s == {(x,z). EX y. (x, y) : s & (y, z) : r}"
- Image :: "[('a * 'b) set, 'a set] => 'b set" (infixl "``" 90)
+definition
+ Image :: "[('a * 'b) set, 'a set] => 'b set"
+ (infixl "``" 90) where
"r `` s == {y. EX x:s. (x,y):r}"
- Id :: "('a * 'a) set" -- {* the identity relation *}
+definition
+ Id :: "('a * 'a) set" where -- {* the identity relation *}
"Id == {p. EX x. p = (x,x)}"
- diag :: "'a set => ('a * 'a) set" -- {* diagonal: identity over a set *}
+definition
+ diag :: "'a set => ('a * 'a) set" where -- {* diagonal: identity over a set *}
"diag A == \<Union>x\<in>A. {(x,x)}"
- Domain :: "('a * 'b) set => 'a set"
+definition
+ Domain :: "('a * 'b) set => 'a set" where
"Domain r == {x. EX y. (x,y):r}"
- Range :: "('a * 'b) set => 'b set"
+definition
+ Range :: "('a * 'b) set => 'b set" where
"Range r == Domain(r^-1)"
- Field :: "('a * 'a) set => 'a set"
+definition
+ Field :: "('a * 'a) set => 'a set" where
"Field r == Domain r \<union> Range r"
- refl :: "['a set, ('a * 'a) set] => bool" -- {* reflexivity over a set *}
+definition
+ refl :: "['a set, ('a * 'a) set] => bool" where -- {* reflexivity over a set *}
"refl A r == r \<subseteq> A \<times> A & (ALL x: A. (x,x) : r)"
- sym :: "('a * 'a) set => bool" -- {* symmetry predicate *}
+definition
+ sym :: "('a * 'a) set => bool" where -- {* symmetry predicate *}
"sym r == ALL x y. (x,y): r --> (y,x): r"
- antisym:: "('a * 'a) set => bool" -- {* antisymmetry predicate *}
+definition
+ antisym :: "('a * 'a) set => bool" where -- {* antisymmetry predicate *}
"antisym r == ALL x y. (x,y):r --> (y,x):r --> x=y"
- trans :: "('a * 'a) set => bool" -- {* transitivity predicate *}
+definition
+ trans :: "('a * 'a) set => bool" where -- {* transitivity predicate *}
"trans r == (ALL x y z. (x,y):r --> (y,z):r --> (x,z):r)"
- single_valued :: "('a * 'b) set => bool"
+definition
+ single_valued :: "('a * 'b) set => bool" where
"single_valued r == ALL x y. (x,y):r --> (ALL z. (x,z):r --> y=z)"
- inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set"
+definition
+ inv_image :: "('b * 'b) set => ('a => 'b) => ('a * 'a) set" where
"inv_image r f == {(x, y). (f x, f y) : r}"
abbreviation
- reflexive :: "('a * 'a) set => bool" -- {* reflexivity over a type *}
+ reflexive :: "('a * 'a) set => bool" where -- {* reflexivity over a type *}
"reflexive == refl UNIV"
--- a/src/HOL/Set.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Set.thy Fri Nov 17 02:20:03 2006 +0100
@@ -38,7 +38,7 @@
"op :" :: "'a => 'a set => bool" -- "membership"
notation
- "op :" ("op :")
+ "op :" ("op :") and
"op :" ("(_/ : _)" [50, 51] 50)
local
@@ -47,32 +47,32 @@
subsection {* Additional concrete syntax *}
abbreviation
- range :: "('a => 'b) => 'b set" -- "of function"
+ range :: "('a => 'b) => 'b set" where -- "of function"
"range f == f ` UNIV"
abbreviation
- "not_mem x A == ~ (x : A)" -- "non-membership"
+ "not_mem x A == ~ (x : A)" -- "non-membership"
notation
- not_mem ("op ~:")
+ not_mem ("op ~:") and
not_mem ("(_/ ~: _)" [50, 51] 50)
notation (xsymbols)
- "op Int" (infixl "\<inter>" 70)
- "op Un" (infixl "\<union>" 65)
- "op :" ("op \<in>")
- "op :" ("(_/ \<in> _)" [50, 51] 50)
- not_mem ("op \<notin>")
- not_mem ("(_/ \<notin> _)" [50, 51] 50)
- Union ("\<Union>_" [90] 90)
+ "op Int" (infixl "\<inter>" 70) and
+ "op Un" (infixl "\<union>" 65) and
+ "op :" ("op \<in>") and
+ "op :" ("(_/ \<in> _)" [50, 51] 50) and
+ not_mem ("op \<notin>") and
+ not_mem ("(_/ \<notin> _)" [50, 51] 50) and
+ Union ("\<Union>_" [90] 90) and
Inter ("\<Inter>_" [90] 90)
notation (HTML output)
- "op Int" (infixl "\<inter>" 70)
- "op Un" (infixl "\<union>" 65)
- "op :" ("op \<in>")
- "op :" ("(_/ \<in> _)" [50, 51] 50)
- not_mem ("op \<notin>")
+ "op Int" (infixl "\<inter>" 70) and
+ "op Un" (infixl "\<union>" 65) and
+ "op :" ("op \<in>") and
+ "op :" ("(_/ \<in> _)" [50, 51] 50) and
+ not_mem ("op \<notin>") and
not_mem ("(_/ \<notin> _)" [50, 51] 50)
syntax
@@ -149,33 +149,37 @@
psubset_def: "A < B == (A::'a set) <= B & ~ A=B" ..
abbreviation
- subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"
+ subset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
"subset == less"
- subset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool"
+
+abbreviation
+ subset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" where
"subset_eq == less_eq"
notation (output)
- subset ("op <")
- subset ("(_/ < _)" [50, 51] 50)
- subset_eq ("op <=")
+ subset ("op <") and
+ subset ("(_/ < _)" [50, 51] 50) and
+ subset_eq ("op <=") and
subset_eq ("(_/ <= _)" [50, 51] 50)
notation (xsymbols)
- subset ("op \<subset>")
- subset ("(_/ \<subset> _)" [50, 51] 50)
- subset_eq ("op \<subseteq>")
+ subset ("op \<subset>") and
+ subset ("(_/ \<subset> _)" [50, 51] 50) and
+ subset_eq ("op \<subseteq>") and
subset_eq ("(_/ \<subseteq> _)" [50, 51] 50)
notation (HTML output)
- subset ("op \<subset>")
- subset ("(_/ \<subset> _)" [50, 51] 50)
- subset_eq ("op \<subseteq>")
+ subset ("op \<subset>") and
+ subset ("(_/ \<subset> _)" [50, 51] 50) and
+ subset_eq ("op \<subseteq>") and
subset_eq ("(_/ \<subseteq> _)" [50, 51] 50)
abbreviation (input)
- supset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "\<supset>" 50)
+ supset :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "\<supset>" 50) where
"supset == greater"
- supset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "\<supseteq>" 50)
+
+abbreviation (input)
+ supset_eq :: "'a set \<Rightarrow> 'a set \<Rightarrow> bool" (infixl "\<supseteq>" 50) where
"supset_eq == greater_eq"
@@ -216,6 +220,7 @@
"\<exists>A\<subseteq>B. P" => "EX A. A \<subseteq> B & P"
"\<exists>!A\<subseteq>B. P" => "EX! A. A \<subseteq> B & P"
+(* FIXME re-use version in Orderings.thy *)
print_translation {*
let
fun
--- a/src/HOL/Transitive_Closure.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Transitive_Closure.thy Fri Nov 17 02:20:03 2006 +0100
@@ -37,17 +37,17 @@
trancl_into_trancl [Pure.intro]: "(a, b) : r^+ ==> (b, c) : r ==> (a,c) : r^+"
abbreviation
- reflcl :: "('a \<times> 'a) set => ('a \<times> 'a) set" ("(_^=)" [1000] 999)
+ reflcl :: "('a \<times> 'a) set => ('a \<times> 'a) set" ("(_^=)" [1000] 999) where
"r^= == r \<union> Id"
notation (xsymbols)
- rtrancl ("(_\<^sup>*)" [1000] 999)
- trancl ("(_\<^sup>+)" [1000] 999)
+ rtrancl ("(_\<^sup>*)" [1000] 999) and
+ trancl ("(_\<^sup>+)" [1000] 999) and
reflcl ("(_\<^sup>=)" [1000] 999)
notation (HTML output)
- rtrancl ("(_\<^sup>*)" [1000] 999)
- trancl ("(_\<^sup>+)" [1000] 999)
+ rtrancl ("(_\<^sup>*)" [1000] 999) and
+ trancl ("(_\<^sup>+)" [1000] 999) and
reflcl ("(_\<^sup>=)" [1000] 999)
--- a/src/HOL/Unix/Unix.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/Unix/Unix.thy Fri Nov 17 02:20:03 2006 +0100
@@ -166,6 +166,7 @@
Val (att, text) \<Rightarrow> att
| Env att dir \<Rightarrow> att)"
+definition
"map_attributes f file =
(case file of
Val (att, text) \<Rightarrow> Val (f att, text)
@@ -830,6 +831,7 @@
[Mkdir user\<^isub>1 perms\<^isub>1 [user\<^isub>1, name\<^isub>1],
Mkdir user\<^isub>2 perms\<^isub>2 [user\<^isub>1, name\<^isub>1, name\<^isub>2],
Creat user\<^isub>2 perms\<^isub>2 [user\<^isub>1, name\<^isub>1, name\<^isub>2, name\<^isub>3]]"
+definition
"bogus_path = [user\<^isub>1, name\<^isub>1, name\<^isub>2]"
text {*
--- a/src/HOL/W0/W0.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/W0/W0.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,7 +12,7 @@
datatype 'a maybe = Ok 'a | Fail
definition
- bind :: "'a maybe \<Rightarrow> ('a \<Rightarrow> 'b maybe) \<Rightarrow> 'b maybe" (infixl "\<bind>" 60)
+ bind :: "'a maybe \<Rightarrow> ('a \<Rightarrow> 'b maybe) \<Rightarrow> 'b maybe" (infixl "\<bind>" 60) where
"m \<bind> f = (case m of Ok r \<Rightarrow> f r | Fail \<Rightarrow> Fail)"
syntax
@@ -85,11 +85,12 @@
"free_tv (x # xs) = free_tv x \<union> free_tv xs"
definition
- dom :: "subst \<Rightarrow> nat set"
+ dom :: "subst \<Rightarrow> nat set" where
"dom s = {n. s n \<noteq> TVar n}"
-- {* domain of a substitution *}
- cod :: "subst \<Rightarrow> nat set"
+definition
+ cod :: "subst \<Rightarrow> nat set" where
"cod s = (\<Union>m \<in> dom s. free_tv (s m))"
-- {* codomain of a substitutions: the introduced variables *}
@@ -103,14 +104,14 @@
*}
definition
- new_tv :: "nat \<Rightarrow> 'a::type_struct \<Rightarrow> bool"
+ new_tv :: "nat \<Rightarrow> 'a::type_struct \<Rightarrow> bool" where
"new_tv n ts = (\<forall>m. m \<in> free_tv ts \<longrightarrow> m < n)"
subsubsection {* Identity substitution *}
definition
- id_subst :: subst
+ id_subst :: subst where
"id_subst = (\<lambda>n. TVar n)"
lemma app_subst_id_te [simp]:
@@ -384,7 +385,7 @@
has_type :: "(typ list \<times> expr \<times> typ) set"
abbreviation
- has_type_rel ("((_) |-/ (_) :: (_))" [60, 0, 60] 60)
+ has_type_rel ("((_) |-/ (_) :: (_))" [60, 0, 60] 60) where
"a |- e :: t == (a, e, t) \<in> has_type"
inductive has_type
--- a/src/HOL/ex/Abstract_NAT.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Abstract_NAT.thy Fri Nov 17 02:20:03 2006 +0100
@@ -66,7 +66,7 @@
text {* \medskip The recursion operator -- polymorphic! *}
definition
- rec :: "'a \<Rightarrow> ('n \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'n \<Rightarrow> 'a"
+ rec :: "'a \<Rightarrow> ('n \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'n \<Rightarrow> 'a" where
"rec e r x = (THE y. Rec e r x y)"
lemma rec_eval:
@@ -92,7 +92,7 @@
text {* \medskip Example: addition (monomorphic) *}
definition
- add :: "'n \<Rightarrow> 'n \<Rightarrow> 'n"
+ add :: "'n \<Rightarrow> 'n \<Rightarrow> 'n" where
"add m n = rec n (\<lambda>_ k. succ k) m"
lemma add_zero [simp]: "add zero n = n"
@@ -116,7 +116,7 @@
text {* \medskip Example: replication (polymorphic) *}
definition
- repl :: "'n \<Rightarrow> 'a \<Rightarrow> 'a list"
+ repl :: "'n \<Rightarrow> 'a \<Rightarrow> 'a list" where
"repl n x = rec [] (\<lambda>_ xs. x # xs) n"
lemma repl_zero [simp]: "repl zero x = []"
--- a/src/HOL/ex/Adder.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Adder.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,7 +15,7 @@
by (cases bv) (simp_all add: bv_to_nat_helper)
definition
- half_adder :: "[bit, bit] => bit list"
+ half_adder :: "[bit, bit] => bit list" where
"half_adder a b = [a bitand b, a bitxor b]"
lemma half_adder_correct: "bv_to_nat (half_adder a b) = bitval a + bitval b"
@@ -28,7 +28,7 @@
by (simp add: half_adder_def)
definition
- full_adder :: "[bit, bit, bit] => bit list"
+ full_adder :: "[bit, bit, bit] => bit list" where
"full_adder a b c =
(let x = a bitxor b in [a bitand b bitor c bitand x, x bitxor c])"
--- a/src/HOL/ex/CTL.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/CTL.thy Fri Nov 17 02:20:03 2006 +0100
@@ -25,7 +25,7 @@
types 'a ctl = "'a set"
definition
- imp :: "'a ctl \<Rightarrow> 'a ctl \<Rightarrow> 'a ctl" (infixr "\<rightarrow>" 75)
+ imp :: "'a ctl \<Rightarrow> 'a ctl \<Rightarrow> 'a ctl" (infixr "\<rightarrow>" 75) where
"p \<rightarrow> q = - p \<union> q"
lemma [intro!]: "p \<inter> p \<rightarrow> q \<subseteq> q" unfolding imp_def by auto
@@ -58,9 +58,11 @@
*}
definition
- EX :: "'a ctl \<Rightarrow> 'a ctl" ("\<EX> _" [80] 90) "\<EX> p = {s. \<exists>s'. (s, s') \<in> \<M> \<and> s' \<in> p}"
- EF :: "'a ctl \<Rightarrow> 'a ctl" ("\<EF> _" [80] 90) "\<EF> p = lfp (\<lambda>s. p \<union> \<EX> s)"
- EG :: "'a ctl \<Rightarrow> 'a ctl" ("\<EG> _" [80] 90) "\<EG> p = gfp (\<lambda>s. p \<inter> \<EX> s)"
+ EX ("\<EX> _" [80] 90) where "\<EX> p = {s. \<exists>s'. (s, s') \<in> \<M> \<and> s' \<in> p}"
+definition
+ EF ("\<EF> _" [80] 90) where "\<EF> p = lfp (\<lambda>s. p \<union> \<EX> s)"
+definition
+ EG ("\<EG> _" [80] 90) where "\<EG> p = gfp (\<lambda>s. p \<inter> \<EX> s)"
text {*
@{text "\<AX>"}, @{text "\<AF>"} and @{text "\<AG>"} are now defined
@@ -69,9 +71,11 @@
*}
definition
- AX :: "'a ctl \<Rightarrow> 'a ctl" ("\<AX> _" [80] 90) "\<AX> p = - \<EX> - p"
- AF :: "'a ctl \<Rightarrow> 'a ctl" ("\<AF> _" [80] 90) "\<AF> p = - \<EG> - p"
- AG :: "'a ctl \<Rightarrow> 'a ctl" ("\<AG> _" [80] 90) "\<AG> p = - \<EF> - p"
+ AX ("\<AX> _" [80] 90) where "\<AX> p = - \<EX> - p"
+definition
+ AF ("\<AF> _" [80] 90) where "\<AF> p = - \<EG> - p"
+definition
+ AG ("\<AG> _" [80] 90) where "\<AG> p = - \<EF> - p"
lemmas [simp] = EX_def EG_def AX_def EF_def AF_def AG_def
--- a/src/HOL/ex/Classpackage.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Classpackage.thy Fri Nov 17 02:20:03 2006 +0100
@@ -97,8 +97,8 @@
qed
definition (in monoid)
- units :: "'a set"
- units_def: "units = { y. \<exists>x. x \<^loc>\<otimes> y = \<^loc>\<one> \<and> y \<^loc>\<otimes> x = \<^loc>\<one> }"
+ units :: "'a set" where
+ "units = { y. \<exists>x. x \<^loc>\<otimes> y = \<^loc>\<one> \<and> y \<^loc>\<otimes> x = \<^loc>\<one> }"
lemma (in monoid) inv_obtain:
assumes ass: "x \<in> units"
@@ -139,11 +139,11 @@
"reduce f g (Suc n) x = f x (reduce f g n x)"
definition (in monoid)
- npow :: "nat \<Rightarrow> 'a \<Rightarrow> 'a"
+ npow :: "nat \<Rightarrow> 'a \<Rightarrow> 'a" where
npow_def_prim: "npow n x = reduce (op \<^loc>\<otimes>) \<^loc>\<one> n x"
abbreviation (in monoid)
- abbrev_npow :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infix "\<^loc>\<up>" 75)
+ abbrev_npow :: "'a \<Rightarrow> nat \<Rightarrow> 'a" (infix "\<^loc>\<up>" 75) where
"x \<^loc>\<up> n \<equiv> npow n x"
lemma (in monoid) npow_def:
@@ -272,12 +272,12 @@
using invr invl by simp
definition (in group)
- pow :: "int \<Rightarrow> 'a \<Rightarrow> 'a"
- pow_def: "pow k x = (if k < 0 then \<^loc>\<div> (monoid.npow (op \<^loc>\<otimes>) \<^loc>\<one> (nat (-k)) x)
+ pow :: "int \<Rightarrow> 'a \<Rightarrow> 'a" where
+ "pow k x = (if k < 0 then \<^loc>\<div> (monoid.npow (op \<^loc>\<otimes>) \<^loc>\<one> (nat (-k)) x)
else (monoid.npow (op \<^loc>\<otimes>) \<^loc>\<one> (nat k) x))"
abbreviation (in group)
- abbrev_pow :: "'a \<Rightarrow> int \<Rightarrow> 'a" (infix "\<^loc>\<up>" 75)
+ abbrev_pow :: "'a \<Rightarrow> int \<Rightarrow> 'a" (infix "\<^loc>\<up>" 75) where
"x \<^loc>\<up> k \<equiv> pow k x"
lemma (in group) int_pow_zero:
@@ -312,12 +312,12 @@
definition
"X a b c = (a \<otimes> \<one> \<otimes> b, a \<otimes> \<one> \<otimes> b, [a, b] \<otimes> \<one> \<otimes> [a, b, c])"
+definition
"Y a b c = (a, \<div> a) \<otimes> \<one> \<otimes> \<div> (b, \<div> c)"
-definition
- "x1 = X (1::nat) 2 3"
- "x2 = X (1::int) 2 3"
- "y2 = Y (1::int) 2 3"
+definition "x1 = X (1::nat) 2 3"
+definition "x2 = X (1::int) 2 3"
+definition "y2 = Y (1::int) 2 3"
code_gen "op \<otimes>" \<one> inv
code_gen X Y (SML) (Haskell)
--- a/src/HOL/ex/CodeCollections.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/CodeCollections.thy Fri Nov 17 02:20:03 2006 +0100
@@ -55,15 +55,19 @@
qed
definition (in ordered)
- min :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+ min :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"min x y = (if x \<^loc><<= y then x else y)"
- max :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+
+definition (in ordered)
+ max :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
"max x y = (if x \<^loc><<= y then y else x)"
definition
- min :: "'a::ordered \<Rightarrow> 'a \<Rightarrow> 'a"
+ min :: "'a::ordered \<Rightarrow> 'a \<Rightarrow> 'a" where
"min x y = (if x <<= y then x else y)"
- max :: "'a::ordered \<Rightarrow> 'a \<Rightarrow> 'a"
+
+definition
+ max :: "'a::ordered \<Rightarrow> 'a \<Rightarrow> 'a" where
"max x y = (if x <<= y then y else x)"
fun abs_sorted :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list \<Rightarrow> bool"
@@ -366,15 +370,15 @@
"get_index p n (x#xs) = (if p x then Some n else get_index p (Suc n) xs)"
definition
- between :: "'a::enum \<Rightarrow> 'a \<Rightarrow> 'a option"
+ between :: "'a::enum \<Rightarrow> 'a \<Rightarrow> 'a option" where
"between x y = get_first (\<lambda>z. x << z & z << y) enum"
definition
- index :: "'a::enum \<Rightarrow> nat"
+ index :: "'a::enum \<Rightarrow> nat" where
"index x = the (get_index (\<lambda>y. y = x) 0 enum)"
definition
- add :: "'a::enum \<Rightarrow> 'a \<Rightarrow> 'a"
+ add :: "'a::enum \<Rightarrow> 'a \<Rightarrow> 'a" where
"add x y =
(let
enm = enum
@@ -387,9 +391,8 @@
"sum [] = inf"
"sum (x#xs) = add x (sum xs)"
-definition
- "test1 = sum [None, Some True, None, Some False]"
- "test2 = (inf :: nat \<times> unit)"
+definition "test1 = sum [None, Some True, None, Some False]"
+definition "test2 = (inf :: nat \<times> unit)"
code_gen "op <<="
code_gen "op <<"
--- a/src/HOL/ex/CodeEmbed.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/CodeEmbed.thy Fri Nov 17 02:20:03 2006 +0100
@@ -22,9 +22,10 @@
| TFix vname sort (infix "\<Colon>\<epsilon>" 117)
abbreviation
- Fun :: "typ \<Rightarrow> typ \<Rightarrow> typ" (infixr "\<rightarrow>" 115)
+ Fun :: "typ \<Rightarrow> typ \<Rightarrow> typ" (infixr "\<rightarrow>" 115) where
"ty1 \<rightarrow> ty2 \<equiv> Type (STR ''fun'') [ty1, ty2]"
- Funs :: "typ list \<Rightarrow> typ \<Rightarrow> typ" (infixr "{\<rightarrow>}" 115)
+abbreviation
+ Funs :: "typ list \<Rightarrow> typ \<Rightarrow> typ" (infixr "{\<rightarrow>}" 115) where
"tys {\<rightarrow>} ty \<equiv> foldr (op \<rightarrow>) tys ty"
datatype "term" =
@@ -33,7 +34,7 @@
| App "term" "term" (infixl "\<bullet>" 110)
abbreviation
- Apps :: "term \<Rightarrow> term list \<Rightarrow> term" (infixl "{\<bullet>}" 110)
+ Apps :: "term \<Rightarrow> term list \<Rightarrow> term" (infixl "{\<bullet>}" 110) where
"t {\<bullet>} ts \<equiv> foldl (op \<bullet>) t ts"
--- a/src/HOL/ex/CodeRandom.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/CodeRandom.thy Fri Nov 17 02:20:03 2006 +0100
@@ -30,7 +30,7 @@
random_seed :: "randseed \<Rightarrow> nat"
definition
- random :: "nat \<Rightarrow> randseed \<Rightarrow> nat \<times> randseed"
+ random :: "nat \<Rightarrow> randseed \<Rightarrow> nat \<times> randseed" where
"random n s = (random_seed s mod n, random_shift s)"
lemma random_bound:
@@ -45,12 +45,13 @@
"snd (random n s) = random_shift s" unfolding random_def by simp
definition
- select :: "'a list \<Rightarrow> randseed \<Rightarrow> 'a \<times> randseed"
+ select :: "'a list \<Rightarrow> randseed \<Rightarrow> 'a \<times> randseed" where
[simp]: "select xs = (do
n \<leftarrow> random (length xs);
return (nth xs n)
done)"
- select_weight :: "(nat \<times> 'a) list \<Rightarrow> randseed \<Rightarrow> 'a \<times> randseed"
+definition
+ select_weight :: "(nat \<times> 'a) list \<Rightarrow> randseed \<Rightarrow> 'a \<times> randseed" where
[simp]: "select_weight xs = (do
n \<leftarrow> random (foldl (op +) 0 (map fst xs));
return (pick xs n)
@@ -123,7 +124,7 @@
qed
definition
- random_int :: "int \<Rightarrow> randseed \<Rightarrow> int * randseed"
+ random_int :: "int \<Rightarrow> randseed \<Rightarrow> int * randseed" where
"random_int k = (do n \<leftarrow> random (nat k); return (int n) done)"
lemma random_nat [code]:
--- a/src/HOL/ex/Codegenerator.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Codegenerator.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,23 +11,27 @@
subsection {* booleans *}
definition
- xor :: "bool \<Rightarrow> bool \<Rightarrow> bool"
+ xor :: "bool \<Rightarrow> bool \<Rightarrow> bool" where
"xor p q = ((p | q) & \<not> (p & q))"
subsection {* natural numbers *}
definition
- n :: nat
+ n :: nat where
"n = 42"
subsection {* pairs *}
definition
- swap :: "'a * 'b \<Rightarrow> 'b * 'a"
+ swap :: "'a * 'b \<Rightarrow> 'b * 'a" where
"swap p = (let (x, y) = p in (y, x))"
- appl :: "('a \<Rightarrow> 'b) * 'a \<Rightarrow> 'b"
+
+definition
+ appl :: "('a \<Rightarrow> 'b) * 'a \<Rightarrow> 'b" where
"appl p = (let (f, x) = p in f x)"
- snd_three :: "'a * 'b * 'c => 'b"
+
+definition
+ snd_three :: "'a * 'b * 'c => 'b" where
"snd_three a = id (\<lambda>(a, b, c). b) a"
lemma [code]:
@@ -41,7 +45,7 @@
subsection {* integers *}
definition
- k :: "int"
+ k :: "int" where
"k = -42"
function
@@ -59,9 +63,11 @@
subsection {* lists *}
definition
- ps :: "nat list"
+ ps :: "nat list" where
"ps = [2, 3, 5, 7, 11]"
- qs :: "nat list"
+
+definition
+ qs :: "nat list" where
"qs == rev ps"
subsection {* mutual datatypes *}
--- a/src/HOL/ex/Higher_Order_Logic.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Higher_Order_Logic.thy Fri Nov 17 02:20:03 2006 +0100
@@ -80,21 +80,31 @@
subsubsection {* Derived connectives *}
definition
- false :: o ("\<bottom>")
+ false :: o ("\<bottom>") where
"\<bottom> \<equiv> \<forall>A. A"
- true :: o ("\<top>")
+
+definition
+ true :: o ("\<top>") where
"\<top> \<equiv> \<bottom> \<longrightarrow> \<bottom>"
- not :: "o \<Rightarrow> o" ("\<not> _" [40] 40)
+
+definition
+ not :: "o \<Rightarrow> o" ("\<not> _" [40] 40) where
"not \<equiv> \<lambda>A. A \<longrightarrow> \<bottom>"
- conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35)
+
+definition
+ conj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<and>" 35) where
"conj \<equiv> \<lambda>A B. \<forall>C. (A \<longrightarrow> B \<longrightarrow> C) \<longrightarrow> C"
- disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30)
+
+definition
+ disj :: "o \<Rightarrow> o \<Rightarrow> o" (infixr "\<or>" 30) where
"disj \<equiv> \<lambda>A B. \<forall>C. (A \<longrightarrow> C) \<longrightarrow> (B \<longrightarrow> C) \<longrightarrow> C"
- Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10)
+
+definition
+ Ex :: "('a \<Rightarrow> o) \<Rightarrow> o" (binder "\<exists>" 10) where
"Ex \<equiv> \<lambda>P. \<forall>C. (\<forall>x. P x \<longrightarrow> C) \<longrightarrow> C"
abbreviation
- not_equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl "\<noteq>" 50)
+ not_equal :: "'a \<Rightarrow> 'a \<Rightarrow> o" (infixl "\<noteq>" 50) where
"x \<noteq> y \<equiv> \<not> (x = y)"
theorem falseE [elim]: "\<bottom> \<Longrightarrow> A"
--- a/src/HOL/ex/InductiveInvariant.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/InductiveInvariant.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,14 +15,14 @@
text "S is an inductive invariant of the functional F with respect to the wellfounded relation r."
definition
- indinv :: "('a * 'a) set => ('a => 'b => bool) => (('a => 'b) => ('a => 'b)) => bool"
+ indinv :: "('a * 'a) set => ('a => 'b => bool) => (('a => 'b) => ('a => 'b)) => bool" where
"indinv r S F = (\<forall>f x. (\<forall>y. (y,x) : r --> S y (f y)) --> S x (F f x))"
text "S is an inductive invariant of the functional F on set D with respect to the wellfounded relation r."
definition
- indinv_on :: "('a * 'a) set => 'a set => ('a => 'b => bool) => (('a => 'b) => ('a => 'b)) => bool"
+ indinv_on :: "('a * 'a) set => 'a set => ('a => 'b => bool) => (('a => 'b) => ('a => 'b)) => bool" where
"indinv_on r D S F = (\<forall>f. \<forall>x\<in>D. (\<forall>y\<in>D. (y,x) \<in> r --> S y (f y)) --> S x (F f x))"
--- a/src/HOL/ex/Lagrange.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Lagrange.thy Fri Nov 17 02:20:03 2006 +0100
@@ -17,7 +17,7 @@
theorem. *}
definition
- sq :: "'a::times => 'a"
+ sq :: "'a::times => 'a" where
"sq x == x*x"
text {* The following lemma essentially shows that every natural
--- a/src/HOL/ex/MonoidGroup.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/MonoidGroup.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,16 +15,18 @@
inv :: "'a => 'a"
definition
- monoid :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |) => bool"
+ monoid :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |) => bool" where
"monoid M = (\<forall>x y z.
times M (times M x y) z = times M x (times M y z) \<and>
times M (one M) x = x \<and> times M x (one M) = x)"
- group :: "(| times :: 'a => 'a => 'a, one :: 'a, inv :: 'a => 'a, ... :: 'b |) => bool"
+definition
+ group :: "(| times :: 'a => 'a => 'a, one :: 'a, inv :: 'a => 'a, ... :: 'b |) => bool" where
"group G = (monoid G \<and> (\<forall>x. times G (inv G x) x = one G))"
+definition
reverse :: "(| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |) =>
- (| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |)"
+ (| times :: 'a => 'a => 'a, one :: 'a, ... :: 'b |)" where
"reverse M = M (| times := \<lambda>x y. times M y x |)"
end
--- a/src/HOL/ex/PER.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/PER.thy Fri Nov 17 02:20:03 2006 +0100
@@ -45,7 +45,7 @@
*}
definition
- "domain" :: "'a::partial_equiv set"
+ "domain" :: "'a::partial_equiv set" where
"domain = {x. x \<sim> x}"
lemma domainI [intro]: "x \<sim> x ==> x \<in> domain"
@@ -165,7 +165,7 @@
*}
definition
- eqv_class :: "('a::partial_equiv) => 'a quot" ("\<lfloor>_\<rfloor>")
+ eqv_class :: "('a::partial_equiv) => 'a quot" ("\<lfloor>_\<rfloor>") where
"\<lfloor>a\<rfloor> = Abs_quot {x. a \<sim> x}"
theorem quot_rep: "\<exists>a. A = \<lfloor>a\<rfloor>"
@@ -232,7 +232,7 @@
subsection {* Picking representing elements *}
definition
- pick :: "'a::partial_equiv quot => 'a"
+ pick :: "'a::partial_equiv quot => 'a" where
"pick A = (SOME a. A = \<lfloor>a\<rfloor>)"
theorem pick_eqv' [intro?, simp]: "a \<in> domain ==> pick \<lfloor>a\<rfloor> \<sim> a"
--- a/src/HOL/ex/Primrec.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Primrec.thy Fri Nov 17 02:20:03 2006 +0100
@@ -43,19 +43,23 @@
text {* The set of primitive recursive functions of type @{typ "nat list => nat"}. *}
definition
- SC :: "nat list => nat"
+ SC :: "nat list => nat" where
"SC l = Suc (zeroHd l)"
- CONSTANT :: "nat => nat list => nat"
+definition
+ CONSTANT :: "nat => nat list => nat" where
"CONSTANT k l = k"
- PROJ :: "nat => nat list => nat"
+definition
+ PROJ :: "nat => nat list => nat" where
"PROJ i l = zeroHd (drop i l)"
- COMP :: "(nat list => nat) => (nat list => nat) list => nat list => nat"
+definition
+ COMP :: "(nat list => nat) => (nat list => nat) list => nat list => nat" where
"COMP g fs l = g (map (\<lambda>f. f l) fs)"
- PREC :: "(nat list => nat) => (nat list => nat) => nat list => nat"
+definition
+ PREC :: "(nat list => nat) => (nat list => nat) => nat list => nat" where
"PREC f g l =
(case l of
[] => 0
--- a/src/HOL/ex/Records.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Records.thy Fri Nov 17 02:20:03 2006 +0100
@@ -51,9 +51,11 @@
subsubsection {* Record selection and record update *}
definition
- getX :: "'a point_scheme => nat"
+ getX :: "'a point_scheme => nat" where
"getX r = xpos r"
- setX :: "'a point_scheme => nat => 'a point_scheme"
+
+definition
+ setX :: "'a point_scheme => nat => 'a point_scheme" where
"setX r n = r (| xpos := n |)"
@@ -145,14 +147,14 @@
*}
definition
- foo5 :: nat
+ foo5 :: nat where
"foo5 = getX (| xpos = 1, ypos = 0 |)"
text {* \medskip Manipulating the ``@{text "..."}'' (more) part. *}
definition
- incX :: "'a point_scheme => 'a point_scheme"
+ incX :: "'a point_scheme => 'a point_scheme" where
"incX r = (| xpos = xpos r + 1, ypos = ypos r, ... = point.more r |)"
lemma "incX r = setX r (Suc (getX r))"
@@ -162,7 +164,7 @@
text {* An alternative definition. *}
definition
- incX' :: "'a point_scheme => 'a point_scheme"
+ incX' :: "'a point_scheme => 'a point_scheme" where
"incX' r = r (| xpos := xpos r + 1 |)"
@@ -194,7 +196,7 @@
*}
definition
- foo10 :: nat
+ foo10 :: nat where
"foo10 = getX (| xpos = 2, ypos = 0, colour = Blue |)"
@@ -206,7 +208,7 @@
*}
definition
- foo11 :: cpoint
+ foo11 :: cpoint where
"foo11 = setX (| xpos = 2, ypos = 0, colour = Blue |) 0"
--- a/src/HOL/ex/Reflected_Presburger.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Reflected_Presburger.thy Fri Nov 17 02:20:03 2006 +0100
@@ -530,7 +530,7 @@
"islintn (n0, t) = False"
definition
- islint :: "intterm \<Rightarrow> bool"
+ islint :: "intterm \<Rightarrow> bool" where
"islint t = islintn(0,t)"
(* And the equivalence to the first definition *)
@@ -730,7 +730,7 @@
(* lin_neg neagtes a linear term *)
definition
- lin_neg :: "intterm \<Rightarrow> intterm"
+ lin_neg :: "intterm \<Rightarrow> intterm" where
"lin_neg i = lin_mul ((-1::int),i)"
(* lin_neg has the semantics of Neg *)
@@ -1625,11 +1625,11 @@
(* Definitions and lemmas about gcd and lcm *)
definition
- lcm :: "nat \<times> nat \<Rightarrow> nat"
+ lcm :: "nat \<times> nat \<Rightarrow> nat" where
"lcm = (\<lambda>(m,n). m*n div gcd(m,n))"
definition
- ilcm :: "int \<Rightarrow> int \<Rightarrow> int"
+ ilcm :: "int \<Rightarrow> int \<Rightarrow> int" where
"ilcm = (\<lambda>i.\<lambda>j. int (lcm(nat(abs i),nat(abs j))))"
(* ilcm_dvd12 are needed later *)
@@ -1879,7 +1879,7 @@
(* unitycoeff expects a quantifier free formula an transforms it to an equivalent formula where the bound variable occurs only with coeffitient 1 or -1 *)
definition
- unitycoeff :: "QF \<Rightarrow> QF"
+ unitycoeff :: "QF \<Rightarrow> QF" where
"unitycoeff p =
(let l = formlcm p;
p' = adjustcoeff (l,p)
@@ -5091,7 +5091,7 @@
(* An implementation of sets trough lists *)
definition
- list_insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list"
+ list_insert :: "'a \<Rightarrow> 'a list \<Rightarrow> 'a list" where
"list_insert x xs = (if x mem xs then xs else x#xs)"
lemma list_insert_set: "set (list_insert x xs) = set (x#xs)"
@@ -5368,7 +5368,7 @@
(* An implementation of cooper's method for both plus/minus/infinity *)
(* unify the formula *)
-definition unify:: "QF \<Rightarrow> (QF \<times> intterm list)"
+definition unify:: "QF \<Rightarrow> (QF \<times> intterm list)" where
"unify p =
(let q = unitycoeff p;
B = list_set(bset q);
@@ -5484,7 +5484,7 @@
qed
(* An implementation of cooper's method *)
definition
- cooper:: "QF \<Rightarrow> QF option"
+ cooper:: "QF \<Rightarrow> QF option" where
"cooper p = lift_un (\<lambda>q. decrvars(explode_minf (unify q))) (linform (nnf p))"
(* cooper eliminates quantifiers *)
@@ -5538,7 +5538,7 @@
(* A decision procedure for Presburger Arithmetics *)
definition
- pa:: "QF \<Rightarrow> QF option"
+ pa:: "QF \<Rightarrow> QF option" where
"pa p \<equiv> lift_un psimpl (qelim(cooper, p))"
lemma psimpl_qfree: "isqfree p \<Longrightarrow> isqfree (psimpl p)"
--- a/src/HOL/ex/Sorting.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Sorting.thy Fri Nov 17 02:20:03 2006 +0100
@@ -25,10 +25,11 @@
definition
- total :: "('a \<Rightarrow> 'a \<Rightarrow> bool) => bool"
+ total :: "('a \<Rightarrow> 'a \<Rightarrow> bool) => bool" where
"total r = (\<forall>x y. r x y | r y x)"
- transf :: "('a \<Rightarrow> 'a \<Rightarrow> bool) => bool"
+definition
+ transf :: "('a \<Rightarrow> 'a \<Rightarrow> bool) => bool" where
"transf f = (\<forall>x y z. f x y & f y z --> f x z)"
--- a/src/HOL/ex/Tarski.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/Tarski.thy Fri Nov 17 02:20:03 2006 +0100
@@ -21,75 +21,88 @@
order :: "('a * 'a) set"
definition
- monotone :: "['a => 'a, 'a set, ('a *'a)set] => bool"
+ 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 = (SOME 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 = (SOME 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}"
definition
- Bot :: "'a potype => 'a"
+ 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 (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))}"
- CLF :: "('a potype * ('a => 'a)) set"
+definition
+ CLF :: "('a potype * ('a => 'a)) set" where
"CLF = (SIGMA cl: CompleteLattice.
{f. f: pset cl -> pset cl & monotone f (pset cl) (order cl)})"
- 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}"
definition
- sublattice :: "('a potype * 'a set)set"
+ sublattice :: "('a potype * 'a set)set" where
"sublattice =
(SIGMA cl: CompleteLattice.
{S. S \<subseteq> pset cl &
(| pset = S, order = induced S (order cl) |): CompleteLattice})"
abbreviation
- sublat :: "['a set, 'a potype] => bool" ("_ <<= _" [51,50]50)
+ sublat :: "['a set, 'a potype] => bool" ("_ <<= _" [51,50]50) where
"S <<= cl == S : sublattice `` {cl}"
definition
- dual :: "'a potype => 'a potype"
+ dual :: "'a potype => 'a potype" where
"dual po = (| pset = pset po, order = converse (order po) |)"
locale (open) PO =
--- a/src/HOL/ex/ThreeDivides.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOL/ex/ThreeDivides.thy Fri Nov 17 02:20:03 2006 +0100
@@ -156,7 +156,7 @@
some number n. *}
definition
- sumdig :: "nat \<Rightarrow> nat"
+ sumdig :: "nat \<Rightarrow> nat" where
"sumdig n = (\<Sum>x < nlen n. n div 10^x mod 10)"
text {* Some properties of these functions follow. *}
--- a/src/HOLCF/FOCUS/Buffer.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/FOCUS/Buffer.thy Fri Nov 17 02:20:03 2006 +0100
@@ -42,47 +42,56 @@
SPECS11 = "SPSF11 set"
definition
-
- BufEq_F :: "SPEC11 \<Rightarrow> SPEC11"
+ BufEq_F :: "SPEC11 \<Rightarrow> SPEC11" where
"BufEq_F B = {f. \<forall>d. f\<cdot>(Md d\<leadsto><>) = <> \<and>
(\<forall>x. \<exists>ff\<in>B. f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x) = d\<leadsto>ff\<cdot>x)}"
- BufEq :: "SPEC11"
+definition
+ BufEq :: "SPEC11" where
"BufEq = gfp BufEq_F"
- BufEq_alt :: "SPEC11"
+definition
+ BufEq_alt :: "SPEC11" where
"BufEq_alt = gfp (\<lambda>B. {f. \<forall>d. f\<cdot>(Md d\<leadsto><> ) = <> \<and>
(\<exists>ff\<in>B. (\<forall>x. f\<cdot>(Md d\<leadsto>\<bullet>\<leadsto>x) = d\<leadsto>ff\<cdot>x))})"
- BufAC_Asm_F :: " (M fstream set) \<Rightarrow> (M fstream set)"
+definition
+ BufAC_Asm_F :: " (M fstream set) \<Rightarrow> (M fstream set)" where
"BufAC_Asm_F A = {s. s = <> \<or>
(\<exists>d x. s = Md d\<leadsto>x \<and> (x = <> \<or> (ft\<cdot>x = Def \<bullet> \<and> (rt\<cdot>x)\<in>A)))}"
- BufAC_Asm :: " (M fstream set)"
+definition
+ BufAC_Asm :: " (M fstream set)" where
"BufAC_Asm = gfp BufAC_Asm_F"
+definition
BufAC_Cmt_F :: "((M fstream * D fstream) set) \<Rightarrow>
- ((M fstream * D fstream) set)"
+ ((M fstream * D fstream) set)" where
"BufAC_Cmt_F C = {(s,t). \<forall>d x.
(s = <> \<longrightarrow> t = <> ) \<and>
(s = Md d\<leadsto><> \<longrightarrow> t = <> ) \<and>
(s = Md d\<leadsto>\<bullet>\<leadsto>x \<longrightarrow> (ft\<cdot>t = Def d \<and> (x,rt\<cdot>t)\<in>C))}"
- BufAC_Cmt :: "((M fstream * D fstream) set)"
+definition
+ BufAC_Cmt :: "((M fstream * D fstream) set)" where
"BufAC_Cmt = gfp BufAC_Cmt_F"
- BufAC :: "SPEC11"
+definition
+ BufAC :: "SPEC11" where
"BufAC = {f. \<forall>x. x\<in>BufAC_Asm \<longrightarrow> (x,f\<cdot>x)\<in>BufAC_Cmt}"
- BufSt_F :: "SPECS11 \<Rightarrow> SPECS11"
+definition
+ BufSt_F :: "SPECS11 \<Rightarrow> SPECS11" where
"BufSt_F H = {h. \<forall>s . h s \<cdot><> = <> \<and>
(\<forall>d x. h \<currency> \<cdot>(Md d\<leadsto>x) = h (Sd d)\<cdot>x \<and>
(\<exists>hh\<in>H. h (Sd d)\<cdot>(\<bullet> \<leadsto>x) = d\<leadsto>(hh \<currency>\<cdot>x)))}"
- BufSt_P :: "SPECS11"
+definition
+ BufSt_P :: "SPECS11" where
"BufSt_P = gfp BufSt_F"
- BufSt :: "SPEC11"
+definition
+ BufSt :: "SPEC11" where
"BufSt = {f. \<exists>h\<in>BufSt_P. f = h \<currency>}"
--- a/src/HOLCF/FOCUS/Fstream.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/FOCUS/Fstream.thy Fri Nov 17 02:20:03 2006 +0100
@@ -17,24 +17,27 @@
types 'a fstream = "'a lift stream"
definition
- fscons :: "'a \<Rightarrow> 'a fstream \<rightarrow> 'a fstream"
+ fscons :: "'a \<Rightarrow> 'a fstream \<rightarrow> 'a fstream" where
"fscons a = (\<Lambda> s. Def a && s)"
- fsfilter :: "'a set \<Rightarrow> 'a fstream \<rightarrow> 'a fstream"
+definition
+ fsfilter :: "'a set \<Rightarrow> 'a fstream \<rightarrow> 'a fstream" where
"fsfilter A = (sfilter\<cdot>(flift2 (\<lambda>x. x\<in>A)))"
abbreviation
- emptystream :: "'a fstream" ("<>")
+ emptystream :: "'a fstream" ("<>") where
"<> == \<bottom>"
- fscons' :: "'a \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_~>_)" [66,65] 65)
+abbreviation
+ fscons' :: "'a \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_~>_)" [66,65] 65) where
"a~>s == fscons a\<cdot>s"
- fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_'(C')_)" [64,63] 63)
+abbreviation
+ fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_'(C')_)" [64,63] 63) where
"A(C)s == fsfilter A\<cdot>s"
notation (xsymbols)
- fscons' ("(_\<leadsto>_)" [66,65] 65)
+ fscons' ("(_\<leadsto>_)" [66,65] 65) and
fsfilter' ("(_\<copyright>_)" [64,63] 63)
--- a/src/HOLCF/FOCUS/Fstreams.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/FOCUS/Fstreams.thy Fri Nov 17 02:20:03 2006 +0100
@@ -13,30 +13,37 @@
types 'a fstream = "('a lift) stream"
definition
- fsingleton :: "'a => 'a fstream" ("<_>" [1000] 999)
+ fsingleton :: "'a => 'a fstream" ("<_>" [1000] 999) where
fsingleton_def2: "fsingleton = (%a. Def a && UU)"
- fsfilter :: "'a set \<Rightarrow> 'a fstream \<rightarrow> 'a fstream"
+definition
+ fsfilter :: "'a set \<Rightarrow> 'a fstream \<rightarrow> 'a fstream" where
"fsfilter A = sfilter\<cdot>(flift2 (\<lambda>x. x\<in>A))"
- fsmap :: "('a => 'b) => 'a fstream -> 'b fstream"
+definition
+ fsmap :: "('a => 'b) => 'a fstream -> 'b fstream" where
"fsmap f = smap$(flift2 f)"
- jth :: "nat => 'a fstream => 'a"
+definition
+ jth :: "nat => 'a fstream => 'a" where
"jth = (%n s. if Fin n < #s then THE a. i_th n s = Def a else arbitrary)"
- first :: "'a fstream => 'a"
+definition
+ first :: "'a fstream => 'a" where
"first = (%s. jth 0 s)"
- last :: "'a fstream => 'a"
+definition
+ last :: "'a fstream => 'a" where
"last = (%s. case #s of Fin n => (if n~=0 then jth (THE k. Suc k = n) s else arbitrary)
| Infty => arbitrary)"
abbreviation
- emptystream :: "'a fstream" ("<>")
+ emptystream :: "'a fstream" ("<>") where
"<> == \<bottom>"
- fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_'(C')_)" [64,63] 63)
+
+abbreviation
+ fsfilter' :: "'a set \<Rightarrow> 'a fstream \<Rightarrow> 'a fstream" ("(_'(C')_)" [64,63] 63) where
"A(C)s == fsfilter A\<cdot>s"
notation (xsymbols)
--- a/src/HOLCF/FOCUS/Stream_adm.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/FOCUS/Stream_adm.thy Fri Nov 17 02:20:03 2006 +0100
@@ -10,18 +10,19 @@
begin
definition
-
- stream_monoP :: "(('a stream) set \<Rightarrow> ('a stream) set) \<Rightarrow> bool"
+ stream_monoP :: "(('a stream) set \<Rightarrow> ('a stream) set) \<Rightarrow> bool" where
"stream_monoP F = (\<exists>Q i. \<forall>P s. Fin i \<le> #s \<longrightarrow>
(s \<in> F P) = (stream_take i\<cdot>s \<in> Q \<and> iterate i\<cdot>rt\<cdot>s \<in> P))"
- stream_antiP :: "(('a stream) set \<Rightarrow> ('a stream) set) \<Rightarrow> bool"
+definition
+ stream_antiP :: "(('a stream) set \<Rightarrow> ('a stream) set) \<Rightarrow> bool" where
"stream_antiP F = (\<forall>P x. \<exists>Q i.
(#x < Fin i \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow> y \<in> F P \<longrightarrow> x \<in> F P)) \<and>
(Fin i <= #x \<longrightarrow> (\<forall>y. x \<sqsubseteq> y \<longrightarrow>
(y \<in> F P) = (stream_take i\<cdot>y \<in> Q \<and> iterate i\<cdot>rt\<cdot>y \<in> P))))"
- antitonP :: "'a set => bool"
+definition
+ antitonP :: "'a set => bool" where
"antitonP P = (\<forall>x y. x \<sqsubseteq> y \<longrightarrow> y\<in>P \<longrightarrow> x\<in>P)"
--- a/src/HOLCF/IMP/Denotational.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/IMP/Denotational.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,7 +11,7 @@
subsection "Definition"
definition
- dlift :: "(('a::type) discr -> 'b::pcpo) => ('a lift -> 'b)"
+ dlift :: "(('a::type) discr -> 'b::pcpo) => ('a lift -> 'b)" where
"dlift f = (LAM x. case x of UU => UU | Def y => f\<cdot>(Discr y))"
consts D :: "com => state discr -> state lift"
--- a/src/HOLCF/IMP/HoareEx.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/IMP/HoareEx.thy Fri Nov 17 02:20:03 2006 +0100
@@ -17,7 +17,7 @@
types assn = "state => bool"
definition
- hoare_valid :: "[assn, com, assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" 50)
+ hoare_valid :: "[assn, com, assn] => bool" ("|= {(1_)}/ (_)/ {(1_)}" 50) where
"|= {A} c {B} = (\<forall>s t. A s \<and> D c $(Discr s) = Def t --> B t)"
lemma WHILE_rule_sound:
--- a/src/HOLCF/ex/Dagstuhl.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/ex/Dagstuhl.thy Fri Nov 17 02:20:03 2006 +0100
@@ -8,9 +8,11 @@
y :: "'a"
definition
- YS :: "'a stream"
+ YS :: "'a stream" where
"YS = fix$(LAM x. y && x)"
- YYS :: "'a stream"
+
+definition
+ YYS :: "'a stream" where
"YYS = fix$(LAM z. y && y && z)"
lemma YS_def2: "YS = y && YS"
--- a/src/HOLCF/ex/Dnat.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/ex/Dnat.thy Fri Nov 17 02:20:03 2006 +0100
@@ -10,7 +10,7 @@
domain dnat = dzero | dsucc (dpred :: dnat)
definition
- iterator :: "dnat -> ('a -> 'a) -> 'a -> 'a"
+ iterator :: "dnat -> ('a -> 'a) -> 'a -> 'a" where
"iterator = fix $ (LAM h n f x.
case n of dzero => x
| dsucc $ m => f $ (h $ m $ f $ x))"
--- a/src/HOLCF/ex/Focus_ex.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/ex/Focus_ex.thy Fri Nov 17 02:20:03 2006 +0100
@@ -111,19 +111,22 @@
Rf :: "('b stream * ('b,'c) tc stream * 'c stream * ('b,'c) tc stream) => bool"
definition
- is_f :: "('b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) => bool"
+ is_f :: "('b stream * ('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) => bool" where
"is_f f = (!i1 i2 o1 o2. f$<i1,i2> = <o1,o2> --> Rf(i1,i2,o1,o2))"
+definition
is_net_g :: "('b stream *('b,'c) tc stream -> 'c stream * ('b,'c) tc stream) =>
- 'b stream => 'c stream => bool"
+ 'b stream => 'c stream => bool" where
"is_net_g f x y == (? z.
<y,z> = f$<x,z> &
(!oy hz. <oy,hz> = f$<x,hz> --> z << hz))"
- is_g :: "('b stream -> 'c stream) => bool"
+definition
+ is_g :: "('b stream -> 'c stream) => bool" where
"is_g g == (? f. is_f f & (!x y. g$x = y --> is_net_g f x y))"
- def_g :: "('b stream -> 'c stream) => bool"
+definition
+ def_g :: "('b stream -> 'c stream) => bool" where
"def_g g == (? f. is_f f & g = (LAM x. cfst$(f$<x,fix$(LAM k. csnd$(f$<x,k>))>)))"
--- a/src/HOLCF/ex/Hoare.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/ex/Hoare.thy Fri Nov 17 02:20:03 2006 +0100
@@ -30,10 +30,11 @@
g :: "'a -> 'a"
definition
- p :: "'a -> 'a"
+ p :: "'a -> 'a" where
"p = fix$(LAM f. LAM x. If b1$x then f$(g$x) else x fi)"
- q :: "'a -> 'a"
+definition
+ q :: "'a -> 'a" where
"q = fix$(LAM f. LAM x. If b1$x orelse b2$x then f$(g$x) else x fi)"
--- a/src/HOLCF/ex/Loop.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/ex/Loop.thy Fri Nov 17 02:20:03 2006 +0100
@@ -10,10 +10,11 @@
begin
definition
- step :: "('a -> tr)->('a -> 'a)->'a->'a"
+ step :: "('a -> tr)->('a -> 'a)->'a->'a" where
"step = (LAM b g x. If b$x then g$x else x fi)"
- while :: "('a -> tr)->('a -> 'a)->'a->'a"
+definition
+ while :: "('a -> tr)->('a -> 'a)->'a->'a" where
"while = (LAM b g. fix$(LAM f x. If b$x then f$(g$x) else x fi))"
(* ------------------------------------------------------------------------- *)
--- a/src/HOLCF/ex/Stream.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/HOLCF/ex/Stream.thy Fri Nov 17 02:20:03 2006 +0100
@@ -12,27 +12,31 @@
domain 'a stream = "&&" (ft::'a) (lazy rt::"'a stream") (infixr 65)
definition
- smap :: "('a \<rightarrow> 'b) \<rightarrow> 'a stream \<rightarrow> 'b stream"
+ smap :: "('a \<rightarrow> 'b) \<rightarrow> 'a stream \<rightarrow> 'b stream" where
"smap = fix\<cdot>(\<Lambda> h f s. case s of x && xs \<Rightarrow> f\<cdot>x && h\<cdot>f\<cdot>xs)"
- sfilter :: "('a \<rightarrow> tr) \<rightarrow> 'a stream \<rightarrow> 'a stream"
+definition
+ sfilter :: "('a \<rightarrow> tr) \<rightarrow> 'a stream \<rightarrow> 'a stream" where
"sfilter = fix\<cdot>(\<Lambda> h p s. case s of x && xs \<Rightarrow>
If p\<cdot>x then x && h\<cdot>p\<cdot>xs else h\<cdot>p\<cdot>xs fi)"
- slen :: "'a stream \<Rightarrow> inat" ("#_" [1000] 1000)
+definition
+ slen :: "'a stream \<Rightarrow> inat" ("#_" [1000] 1000) where
"#s = (if stream_finite s then Fin (LEAST n. stream_take n\<cdot>s = s) else \<infinity>)"
(* concatenation *)
definition
- i_rt :: "nat => 'a stream => 'a stream" (* chops the first i elements *)
+ i_rt :: "nat => 'a stream => 'a stream" where (* chops the first i elements *)
"i_rt = (%i s. iterate i$rt$s)"
- i_th :: "nat => 'a stream => 'a" (* the i-th element *)
+definition
+ i_th :: "nat => 'a stream => 'a" where (* the i-th element *)
"i_th = (%i s. ft$(i_rt i s))"
- sconc :: "'a stream => 'a stream => 'a stream" (infixr "ooo" 65)
+definition
+ sconc :: "'a stream => 'a stream => 'a stream" (infixr "ooo" 65) where
"s1 ooo s2 = (case #s1 of
Fin n \<Rightarrow> (SOME s. (stream_take n$s=s1) & (i_rt n s = s2))
| \<infinity> \<Rightarrow> s1)"
@@ -45,7 +49,7 @@
constr_sconc' n (rt$s1) s2"
definition
- constr_sconc :: "'a stream => 'a stream => 'a stream" (* constructive *)
+ constr_sconc :: "'a stream => 'a stream => 'a stream" where (* constructive *)
"constr_sconc s1 s2 = (case #s1 of
Fin n \<Rightarrow> constr_sconc' n s1 s2
| \<infinity> \<Rightarrow> s1)"
--- a/src/ZF/Constructible/AC_in_L.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/AC_in_L.thy Fri Nov 17 02:20:03 2006 +0100
@@ -223,18 +223,20 @@
"DPow(A)"}, we take the minimum such ordinal.*}
definition
- env_form_r :: "[i,i,i]=>i"
+ env_form_r :: "[i,i,i]=>i" where
--{*wellordering on (environment, formula) pairs*}
"env_form_r(f,r,A) ==
rmult(list(A), rlist(A, r),
formula, measure(formula, enum(f)))"
- env_form_map :: "[i,i,i,i]=>i"
+definition
+ env_form_map :: "[i,i,i,i]=>i" where
--{*map from (environment, formula) pairs to ordinals*}
"env_form_map(f,r,A,z)
== ordermap(list(A) * formula, env_form_r(f,r,A)) ` z"
- DPow_ord :: "[i,i,i,i,i]=>o"
+definition
+ DPow_ord :: "[i,i,i,i,i]=>o" where
--{*predicate that holds if @{term k} is a valid index for @{term X}*}
"DPow_ord(f,r,A,X,k) ==
\<exists>env \<in> list(A). \<exists>p \<in> formula.
@@ -242,11 +244,13 @@
X = {x\<in>A. sats(A, p, Cons(x,env))} &
env_form_map(f,r,A,<env,p>) = k"
- DPow_least :: "[i,i,i,i]=>i"
+definition
+ DPow_least :: "[i,i,i,i]=>i" where
--{*function yielding the smallest index for @{term X}*}
"DPow_least(f,r,A,X) == \<mu> k. DPow_ord(f,r,A,X,k)"
- DPow_r :: "[i,i,i]=>i"
+definition
+ DPow_r :: "[i,i,i]=>i" where
--{*a wellordering on @{term "DPow(A)"}*}
"DPow_r(f,r,A) == measure(DPow(A), DPow_least(f,r,A))"
@@ -318,7 +322,7 @@
of wellorderings for smaller ordinals.*}
definition
- rlimit :: "[i,i=>i]=>i"
+ rlimit :: "[i,i=>i]=>i" where
--{*Expresses the wellordering at limit ordinals. The conditional
lets us remove the premise @{term "Limit(i)"} from some theorems.*}
"rlimit(i,r) ==
@@ -329,7 +333,8 @@
(lrank(x') = lrank(x) & <x',x> \<in> r(succ(lrank(x)))))}
else 0"
- Lset_new :: "i=>i"
+definition
+ Lset_new :: "i=>i" where
--{*This constant denotes the set of elements introduced at level
@{term "succ(i)"}*}
"Lset_new(i) == {x \<in> Lset(succ(i)). lrank(x) = i}"
@@ -401,7 +406,7 @@
subsection{*Transfinite Definition of the Wellordering on @{term "L"}*}
definition
- L_r :: "[i, i] => i"
+ L_r :: "[i, i] => i" where
"L_r(f) == %i.
transrec3(i, 0, \<lambda>x r. DPow_r(f, r, Lset(x)),
\<lambda>x r. rlimit(x, \<lambda>y. r`y))"
--- a/src/ZF/Constructible/DPow_absolute.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/DPow_absolute.thy Fri Nov 17 02:20:03 2006 +0100
@@ -23,7 +23,8 @@
successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)"
*)
-definition formula_rec_fm :: "[i, i, i]=>i"
+definition
+ formula_rec_fm :: "[i, i, i]=>i" where
"formula_rec_fm(mh,p,z) ==
Exists(Exists(Exists(
And(finite_ordinal_fm(2),
@@ -80,7 +81,8 @@
subsubsection{*The Operator @{term is_satisfies}*}
(* is_satisfies(M,A,p,z) == is_formula_rec (M, satisfies_MH(M,A), p, z) *)
-definition satisfies_fm :: "[i,i,i]=>i"
+definition
+ satisfies_fm :: "[i,i,i]=>i" where
"satisfies_fm(x) == formula_rec_fm (satisfies_MH_fm(x#+5#+6, 2, 1, 0))"
lemma is_satisfies_type [TC]:
@@ -120,7 +122,7 @@
text{*Relativize the use of @{term sats} within @{term DPow'}
(the comprehension).*}
definition
- is_DPow_sats :: "[i=>o,i,i,i,i] => o"
+ is_DPow_sats :: "[i=>o,i,i,i,i] => o" where
"is_DPow_sats(M,A,env,p,x) ==
\<forall>n1[M]. \<forall>e[M]. \<forall>sp[M].
is_satisfies(M,A,p,sp) --> is_Cons(M,x,env,e) -->
@@ -148,8 +150,9 @@
is_satisfies(M,A,p,sp) --> is_Cons(M,x,env,e) -->
fun_apply(M, sp, e, n1) --> number1(M, n1) *)
-definition DPow_sats_fm :: "[i,i,i,i]=>i"
- "DPow_sats_fm(A,env,p,x) ==
+definition
+ DPow_sats_fm :: "[i,i,i,i]=>i" where
+ "DPow_sats_fm(A,env,p,x) ==
Forall(Forall(Forall(
Implies(satisfies_fm(A#+3,p#+3,0),
Implies(Cons_fm(x#+3,env#+3,1),
@@ -219,7 +222,7 @@
text{*Relativization of the Operator @{term DPow'}*}
definition
- is_DPow' :: "[i=>o,i,i] => o"
+ is_DPow' :: "[i=>o,i,i] => o" where
"is_DPow'(M,A,Z) ==
\<forall>X[M]. X \<in> Z <->
subset(M,X,A) &
@@ -310,7 +313,8 @@
(* is_Collect :: "[i=>o,i,i=>o,i] => o"
"is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z <-> x \<in> A & P(x)" *)
-definition Collect_fm :: "[i, i, i]=>i"
+definition
+ Collect_fm :: "[i, i, i]=>i" where
"Collect_fm(A,is_P,z) ==
Forall(Iff(Member(0,succ(z)),
And(Member(0,succ(A)), is_P)))"
@@ -360,8 +364,9 @@
(* is_Replace :: "[i=>o,i,[i,i]=>o,i] => o"
"is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & P(x,u))" *)
-definition Replace_fm :: "[i, i, i]=>i"
- "Replace_fm(A,is_P,z) ==
+definition
+ Replace_fm :: "[i, i, i]=>i" where
+ "Replace_fm(A,is_P,z) ==
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,A#+2), is_P))))"
@@ -413,7 +418,8 @@
(\<exists>env[M]. \<exists>p[M]. mem_formula(M,p) & mem_list(M,A,env) &
is_Collect(M, A, is_DPow_sats(M,A,env,p), X))" *)
-definition DPow'_fm :: "[i,i]=>i"
+definition
+ DPow'_fm :: "[i,i]=>i" where
"DPow'_fm(A,Z) ==
Forall(
Iff(Member(0,succ(Z)),
@@ -452,7 +458,7 @@
subsection{*A Locale for Relativizing the Operator @{term Lset}*}
definition
- transrec_body :: "[i=>o,i,i,i,i] => o"
+ transrec_body :: "[i=>o,i,i,i,i] => o" where
"transrec_body(M,g,x) ==
\<lambda>y z. \<exists>gy[M]. y \<in> x & fun_apply(M,g,y,gy) & is_DPow'(M,gy,z)"
@@ -504,7 +510,7 @@
text{*Relativization of the Operator @{term Lset}*}
definition
- is_Lset :: "[i=>o, i, i] => o"
+ is_Lset :: "[i=>o, i, i] => o" where
--{*We can use the term language below because @{term is_Lset} will
not have to be internalized: it isn't used in any instance of
separation.*}
@@ -610,7 +616,7 @@
definition
- constructible :: "[i=>o,i] => o"
+ constructible :: "[i=>o,i] => o" where
"constructible(M,x) ==
\<exists>i[M]. \<exists>Li[M]. ordinal(M,i) & is_Lset(M,i,Li) & x \<in> Li"
--- a/src/ZF/Constructible/Datatype_absolute.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/Datatype_absolute.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,10 +11,11 @@
subsection{*The lfp of a continuous function can be expressed as a union*}
definition
- directed :: "i=>o"
+ directed :: "i=>o" where
"directed(A) == A\<noteq>0 & (\<forall>x\<in>A. \<forall>y\<in>A. x \<union> y \<in> A)"
- contin :: "(i=>i) => o"
+definition
+ contin :: "(i=>i) => o" where
"contin(h) == (\<forall>A. directed(A) --> h(\<Union>A) = (\<Union>X\<in>A. h(X)))"
lemma bnd_mono_iterates_subset: "[|bnd_mono(D, h); n \<in> nat|] ==> h^n (0) <= D"
@@ -114,18 +115,19 @@
subsection {*Absoluteness for "Iterates"*}
definition
-
- iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o"
+ iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" where
"iterates_MH(M,isF,v,n,g,z) ==
is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) & isF(gm,u),
n, z)"
- is_iterates :: "[i=>o, [i,i]=>o, i, i, i] => o"
+definition
+ is_iterates :: "[i=>o, [i,i]=>o, i, i, i] => o" where
"is_iterates(M,isF,v,n,Z) ==
\<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) & membership(M,sn,msn) &
is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)"
- iterates_replacement :: "[i=>o, [i,i]=>o, i] => o"
+definition
+ iterates_replacement :: "[i=>o, [i,i]=>o, i] => o" where
"iterates_replacement(M,isF,v) ==
\<forall>n[M]. n\<in>nat -->
wfrec_replacement(M, iterates_MH(M,isF,v), Memrel(succ(n)))"
@@ -208,7 +210,7 @@
definition
- is_list_functor :: "[i=>o,i,i,i] => o"
+ is_list_functor :: "[i=>o,i,i,i] => o" where
"is_list_functor(M,A,X,Z) ==
\<exists>n1[M]. \<exists>AX[M].
number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)"
@@ -261,7 +263,7 @@
definition
- is_formula_functor :: "[i=>o,i,i] => o"
+ is_formula_functor :: "[i=>o,i,i] => o" where
"is_formula_functor(M,X,Z) ==
\<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M].
omega(M,nat') & cartprod(M,nat',nat',natnat) &
@@ -279,7 +281,7 @@
subsection{*@{term M} Contains the List and Formula Datatypes*}
definition
- list_N :: "[i,i] => i"
+ list_N :: "[i,i] => i" where
"list_N(A,n) == (\<lambda>X. {0} + A * X)^n (0)"
lemma Nil_in_list_N [simp]: "[] \<in> list_N(A,succ(n))"
@@ -340,17 +342,19 @@
done
definition
- is_list_N :: "[i=>o,i,i,i] => o"
+ is_list_N :: "[i=>o,i,i,i] => o" where
"is_list_N(M,A,n,Z) ==
\<exists>zero[M]. empty(M,zero) &
is_iterates(M, is_list_functor(M,A), zero, n, Z)"
-
- mem_list :: "[i=>o,i,i] => o"
+
+definition
+ mem_list :: "[i=>o,i,i] => o" where
"mem_list(M,A,l) ==
\<exists>n[M]. \<exists>listn[M].
finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \<in> listn"
- is_list :: "[i=>o,i,i] => o"
+definition
+ is_list :: "[i=>o,i,i] => o" where
"is_list(M,A,Z) == \<forall>l[M]. l \<in> Z <-> mem_list(M,A,l)"
subsubsection{*Towards Absoluteness of @{term formula_rec}*}
@@ -367,7 +371,7 @@
definition
- formula_N :: "i => i"
+ formula_N :: "i => i" where
"formula_N(n) == (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)"
lemma Member_in_formula_N [simp]:
@@ -442,20 +446,20 @@
done
definition
- is_formula_N :: "[i=>o,i,i] => o"
+ is_formula_N :: "[i=>o,i,i] => o" where
"is_formula_N(M,n,Z) ==
\<exists>zero[M]. empty(M,zero) &
is_iterates(M, is_formula_functor(M), zero, n, Z)"
-definition
-
- mem_formula :: "[i=>o,i] => o"
+definition
+ mem_formula :: "[i=>o,i] => o" where
"mem_formula(M,p) ==
\<exists>n[M]. \<exists>formn[M].
finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \<in> formn"
- is_formula :: "[i=>o,i] => o"
+definition
+ is_formula :: "[i=>o,i] => o" where
"is_formula(M,Z) == \<forall>p[M]. p \<in> Z <-> mem_formula(M,p)"
locale M_datatypes = M_trancl +
@@ -585,15 +589,17 @@
done
definition
- is_eclose_n :: "[i=>o,i,i,i] => o"
+ is_eclose_n :: "[i=>o,i,i,i] => o" where
"is_eclose_n(M,A,n,Z) == is_iterates(M, big_union(M), A, n, Z)"
- mem_eclose :: "[i=>o,i,i] => o"
+definition
+ mem_eclose :: "[i=>o,i,i] => o" where
"mem_eclose(M,A,l) ==
\<exists>n[M]. \<exists>eclosen[M].
finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \<in> eclosen"
- is_eclose :: "[i=>o,i,i] => o"
+definition
+ is_eclose :: "[i=>o,i,i] => o" where
"is_eclose(M,A,Z) == \<forall>u[M]. u \<in> Z <-> mem_eclose(M,A,u)"
@@ -643,15 +649,16 @@
subsection {*Absoluteness for @{term transrec}*}
text{* @{term "transrec(a,H) \<equiv> wfrec(Memrel(eclose({a})), a, H)"} *}
+
definition
-
- is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o"
+ is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" where
"is_transrec(M,MH,a,z) ==
\<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M].
upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
is_wfrec(M,MH,mesa,a,z)"
- transrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o"
+definition
+ transrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where
"transrec_replacement(M,MH,a) ==
\<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M].
upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
@@ -692,7 +699,7 @@
text{*But it is never used.*}
definition
- is_length :: "[i=>o,i,i,i] => o"
+ is_length :: "[i=>o,i,i,i] => o" where
"is_length(M,A,l,n) ==
\<exists>sn[M]. \<exists>list_n[M]. \<exists>list_sn[M].
is_list_N(M,A,n,list_n) & l \<notin> list_n &
@@ -740,7 +747,7 @@
done
definition
- is_nth :: "[i=>o,i,i,i] => o"
+ is_nth :: "[i=>o,i,i,i] => o" where
"is_nth(M,n,l,Z) ==
\<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)"
@@ -758,7 +765,7 @@
subsection{*Relativization and Absoluteness for the @{term formula} Constructors*}
definition
- is_Member :: "[i=>o,i,i,i] => o"
+ is_Member :: "[i=>o,i,i,i] => o" where
--{* because @{term "Member(x,y) \<equiv> Inl(Inl(\<langle>x,y\<rangle>))"}*}
"is_Member(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)"
@@ -772,7 +779,7 @@
by (simp add: Member_def)
definition
- is_Equal :: "[i=>o,i,i,i] => o"
+ is_Equal :: "[i=>o,i,i,i] => o" where
--{* because @{term "Equal(x,y) \<equiv> Inl(Inr(\<langle>x,y\<rangle>))"}*}
"is_Equal(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)"
@@ -785,7 +792,7 @@
by (simp add: Equal_def)
definition
- is_Nand :: "[i=>o,i,i,i] => o"
+ is_Nand :: "[i=>o,i,i,i] => o" where
--{* because @{term "Nand(x,y) \<equiv> Inr(Inl(\<langle>x,y\<rangle>))"}*}
"is_Nand(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)"
@@ -798,7 +805,7 @@
by (simp add: Nand_def)
definition
- is_Forall :: "[i=>o,i,i] => o"
+ is_Forall :: "[i=>o,i,i] => o" where
--{* because @{term "Forall(x) \<equiv> Inr(Inr(p))"}*}
"is_Forall(M,p,Z) == \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)"
@@ -814,8 +821,7 @@
subsection {*Absoluteness for @{term formula_rec}*}
definition
-
- formula_rec_case :: "[[i,i]=>i, [i,i]=>i, [i,i,i,i]=>i, [i,i]=>i, i, i] => i"
+ formula_rec_case :: "[[i,i]=>i, [i,i]=>i, [i,i,i,i]=>i, [i,i]=>i, i, i] => i" where
--{* the instance of @{term formula_case} in @{term formula_rec}*}
"formula_rec_case(a,b,c,d,h) ==
formula_case (a, b,
@@ -847,9 +853,9 @@
subsubsection{*Absoluteness for the Formula Operator @{term depth}*}
+
definition
-
- is_depth :: "[i=>o,i,i] => o"
+ is_depth :: "[i=>o,i,i] => o" where
"is_depth(M,p,n) ==
\<exists>sn[M]. \<exists>formula_n[M]. \<exists>formula_sn[M].
is_formula_N(M,n,formula_n) & p \<notin> formula_n &
@@ -874,9 +880,8 @@
subsubsection{*@{term is_formula_case}: relativization of @{term formula_case}*}
definition
-
is_formula_case ::
- "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o"
+ "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" where
--{*no constraint on non-formulas*}
"is_formula_case(M, is_a, is_b, is_c, is_d, p, z) ==
(\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) --> finite_ordinal(M,y) -->
@@ -910,7 +915,7 @@
subsubsection {*Absoluteness for @{term formula_rec}: Final Results*}
definition
- is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o"
+ is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" where
--{* predicate to relativize the functional @{term formula_rec}*}
"is_formula_rec(M,MH,p,z) ==
\<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) &
--- a/src/ZF/Constructible/Formula.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/Formula.thy Fri Nov 17 02:20:03 2006 +0100
@@ -21,23 +21,29 @@
declare formula.intros [TC]
-definition Neg :: "i=>i"
- "Neg(p) == Nand(p,p)"
+definition
+ Neg :: "i=>i" where
+ "Neg(p) == Nand(p,p)"
-definition And :: "[i,i]=>i"
- "And(p,q) == Neg(Nand(p,q))"
+definition
+ And :: "[i,i]=>i" where
+ "And(p,q) == Neg(Nand(p,q))"
-definition Or :: "[i,i]=>i"
- "Or(p,q) == Nand(Neg(p),Neg(q))"
+definition
+ Or :: "[i,i]=>i" where
+ "Or(p,q) == Nand(Neg(p),Neg(q))"
-definition Implies :: "[i,i]=>i"
- "Implies(p,q) == Nand(p,Neg(q))"
+definition
+ Implies :: "[i,i]=>i" where
+ "Implies(p,q) == Nand(p,Neg(q))"
-definition Iff :: "[i,i]=>i"
- "Iff(p,q) == And(Implies(p,q), Implies(q,p))"
+definition
+ Iff :: "[i,i]=>i" where
+ "Iff(p,q) == And(Implies(p,q), Implies(q,p))"
-definition Exists :: "i=>i"
- "Exists(p) == Neg(Forall(Neg(p)))";
+definition
+ Exists :: "i=>i" where
+ "Exists(p) == Neg(Forall(Neg(p)))";
lemma Neg_type [TC]: "p \<in> formula ==> Neg(p) \<in> formula"
by (simp add: Neg_def)
@@ -79,7 +85,7 @@
by (induct set: formula) simp_all
abbreviation
- sats :: "[i,i,i] => o"
+ sats :: "[i,i,i] => o" where
"sats(A,p,env) == satisfies(A,p)`env = 1"
lemma [simp]:
@@ -246,8 +252,9 @@
subsection{*Renaming Some de Bruijn Variables*}
-definition incr_var :: "[i,i]=>i"
- "incr_var(x,nq) == if x<nq then x else succ(x)"
+definition
+ incr_var :: "[i,i]=>i" where
+ "incr_var(x,nq) == if x<nq then x else succ(x)"
lemma incr_var_lt: "x<nq ==> incr_var(x,nq) = x"
by (simp add: incr_var_def)
@@ -334,8 +341,9 @@
subsection{*Renaming all but the First de Bruijn Variable*}
-definition incr_bv1 :: "i => i"
- "incr_bv1(p) == incr_bv(p)`1"
+definition
+ incr_bv1 :: "i => i" where
+ "incr_bv1(p) == incr_bv(p)`1"
lemma incr_bv1_type [TC]: "p \<in> formula ==> incr_bv1(p) \<in> formula"
@@ -385,7 +393,8 @@
subsection{*Definable Powerset*}
text{*The definable powerset operation: Kunen's definition VI 1.1, page 165.*}
-definition DPow :: "i => i"
+definition
+ DPow :: "i => i" where
"DPow(A) == {X \<in> Pow(A).
\<exists>env \<in> list(A). \<exists>p \<in> formula.
arity(p) \<le> succ(length(env)) &
@@ -507,8 +516,9 @@
subsubsection{*The subset relation*}
-definition subset_fm :: "[i,i]=>i"
- "subset_fm(x,y) == Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))"
+definition
+ subset_fm :: "[i,i]=>i" where
+ "subset_fm(x,y) == Forall(Implies(Member(0,succ(x)), Member(0,succ(y))))"
lemma subset_type [TC]: "[| x \<in> nat; y \<in> nat |] ==> subset_fm(x,y) \<in> formula"
by (simp add: subset_fm_def)
@@ -527,8 +537,9 @@
subsubsection{*Transitive sets*}
-definition transset_fm :: "i=>i"
- "transset_fm(x) == Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))"
+definition
+ transset_fm :: "i=>i" where
+ "transset_fm(x) == Forall(Implies(Member(0,succ(x)), subset_fm(0,succ(x))))"
lemma transset_type [TC]: "x \<in> nat ==> transset_fm(x) \<in> formula"
by (simp add: transset_fm_def)
@@ -547,9 +558,10 @@
subsubsection{*Ordinals*}
-definition ordinal_fm :: "i=>i"
- "ordinal_fm(x) ==
- And(transset_fm(x), Forall(Implies(Member(0,succ(x)), transset_fm(0))))"
+definition
+ ordinal_fm :: "i=>i" where
+ "ordinal_fm(x) ==
+ And(transset_fm(x), Forall(Implies(Member(0,succ(x)), transset_fm(0))))"
lemma ordinal_type [TC]: "x \<in> nat ==> ordinal_fm(x) \<in> formula"
by (simp add: ordinal_fm_def)
@@ -579,11 +591,12 @@
subsection{* Constant Lset: Levels of the Constructible Universe *}
definition
- Lset :: "i=>i"
- "Lset(i) == transrec(i, %x f. \<Union>y\<in>x. DPow(f`y))"
+ Lset :: "i=>i" where
+ "Lset(i) == transrec(i, %x f. \<Union>y\<in>x. DPow(f`y))"
- L :: "i=>o" --{*Kunen's definition VI 1.5, page 167*}
- "L(x) == \<exists>i. Ord(i) & x \<in> Lset(i)"
+definition
+ L :: "i=>o" where --{*Kunen's definition VI 1.5, page 167*}
+ "L(x) == \<exists>i. Ord(i) & x \<in> Lset(i)"
text{*NOT SUITABLE FOR REWRITING -- RECURSIVE!*}
lemma Lset: "Lset(i) = (UN j:i. DPow(Lset(j)))"
@@ -825,8 +838,8 @@
text{*The rank function for the constructible universe*}
definition
- lrank :: "i=>i" --{*Kunen's definition VI 1.7*}
- "lrank(x) == \<mu> i. x \<in> Lset(succ(i))"
+ lrank :: "i=>i" where --{*Kunen's definition VI 1.7*}
+ "lrank(x) == \<mu> i. x \<in> Lset(succ(i))"
lemma L_I: "[|x \<in> Lset(i); Ord(i)|] ==> L(x)"
by (simp add: L_def, blast)
@@ -984,7 +997,8 @@
text{*A simpler version of @{term DPow}: no arity check!*}
-definition DPow' :: "i => i"
+definition
+ DPow' :: "i => i" where
"DPow'(A) == {X \<in> Pow(A).
\<exists>env \<in> list(A). \<exists>p \<in> formula.
X = {x\<in>A. sats(A, p, Cons(x,env))}}"
--- a/src/ZF/Constructible/Internalize.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/Internalize.thy Fri Nov 17 02:20:03 2006 +0100
@@ -10,7 +10,8 @@
subsubsection{*The Formula @{term is_Inl}, Internalized*}
(* is_Inl(M,a,z) == \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z) *)
-definition Inl_fm :: "[i,i]=>i"
+definition
+ Inl_fm :: "[i,i]=>i" where
"Inl_fm(a,z) == Exists(And(empty_fm(0), pair_fm(0,succ(a),succ(z))))"
lemma Inl_type [TC]:
@@ -39,7 +40,8 @@
subsubsection{*The Formula @{term is_Inr}, Internalized*}
(* is_Inr(M,a,z) == \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z) *)
-definition Inr_fm :: "[i,i]=>i"
+definition
+ Inr_fm :: "[i,i]=>i" where
"Inr_fm(a,z) == Exists(And(number1_fm(0), pair_fm(0,succ(a),succ(z))))"
lemma Inr_type [TC]:
@@ -69,7 +71,8 @@
(* is_Nil(M,xs) == \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs) *)
-definition Nil_fm :: "i=>i"
+definition
+ Nil_fm :: "i=>i" where
"Nil_fm(x) == Exists(And(empty_fm(0), Inl_fm(0,succ(x))))"
lemma Nil_type [TC]: "x \<in> nat ==> Nil_fm(x) \<in> formula"
@@ -97,7 +100,8 @@
(* "is_Cons(M,a,l,Z) == \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)" *)
-definition Cons_fm :: "[i,i,i]=>i"
+definition
+ Cons_fm :: "[i,i,i]=>i" where
"Cons_fm(a,l,Z) ==
Exists(And(pair_fm(succ(a),succ(l),0), Inr_fm(0,succ(Z))))"
@@ -128,7 +132,8 @@
(* is_quasilist(M,xs) == is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))" *)
-definition quasilist_fm :: "i=>i"
+definition
+ quasilist_fm :: "i=>i" where
"quasilist_fm(x) ==
Or(Nil_fm(x), Exists(Exists(Cons_fm(1,0,succ(succ(x))))))"
@@ -162,7 +167,8 @@
(is_Nil(M,xs) --> empty(M,H)) &
(\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | H=x) &
(is_quasilist(M,xs) | empty(M,H))" *)
-definition hd_fm :: "[i,i]=>i"
+definition
+ hd_fm :: "[i,i]=>i" where
"hd_fm(xs,H) ==
And(Implies(Nil_fm(xs), empty_fm(H)),
And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(H#+2,1)))),
@@ -198,7 +204,8 @@
(is_Nil(M,xs) --> T=xs) &
(\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | T=l) &
(is_quasilist(M,xs) | empty(M,T))" *)
-definition tl_fm :: "[i,i]=>i"
+definition
+ tl_fm :: "[i,i]=>i" where
"tl_fm(xs,T) ==
And(Implies(Nil_fm(xs), Equal(T,xs)),
And(Forall(Forall(Or(Neg(Cons_fm(1,0,xs#+2)), Equal(T#+2,0)))),
@@ -234,8 +241,9 @@
"is_bool_of_o(M,P,z) == (P & number1(M,z)) | (~P & empty(M,z))" *)
text{*The formula @{term p} has no free variables.*}
-definition bool_of_o_fm :: "[i, i]=>i"
- "bool_of_o_fm(p,z) ==
+definition
+ bool_of_o_fm :: "[i, i]=>i" where
+ "bool_of_o_fm(p,z) ==
Or(And(p,number1_fm(z)),
And(Neg(p),empty_fm(z)))"
@@ -276,8 +284,9 @@
"is_lambda(M, A, is_b, z) ==
\<forall>p[M]. p \<in> z <->
(\<exists>u[M]. \<exists>v[M]. u\<in>A & pair(M,u,v,p) & is_b(u,v))" *)
-definition lambda_fm :: "[i, i, i]=>i"
- "lambda_fm(p,A,z) ==
+definition
+ lambda_fm :: "[i, i, i]=>i" where
+ "lambda_fm(p,A,z) ==
Forall(Iff(Member(0,succ(z)),
Exists(Exists(And(Member(1,A#+3),
And(pair_fm(1,0,2), p))))))"
@@ -315,7 +324,8 @@
(* "is_Member(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" *)
-definition Member_fm :: "[i,i,i]=>i"
+definition
+ Member_fm :: "[i,i,i]=>i" where
"Member_fm(x,y,Z) ==
Exists(Exists(And(pair_fm(x#+2,y#+2,1),
And(Inl_fm(1,0), Inl_fm(0,Z#+2)))))"
@@ -347,7 +357,8 @@
(* "is_Equal(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" *)
-definition Equal_fm :: "[i,i,i]=>i"
+definition
+ Equal_fm :: "[i,i,i]=>i" where
"Equal_fm(x,y,Z) ==
Exists(Exists(And(pair_fm(x#+2,y#+2,1),
And(Inr_fm(1,0), Inl_fm(0,Z#+2)))))"
@@ -379,7 +390,8 @@
(* "is_Nand(M,x,y,Z) ==
\<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" *)
-definition Nand_fm :: "[i,i,i]=>i"
+definition
+ Nand_fm :: "[i,i,i]=>i" where
"Nand_fm(x,y,Z) ==
Exists(Exists(And(pair_fm(x#+2,y#+2,1),
And(Inl_fm(1,0), Inr_fm(0,Z#+2)))))"
@@ -410,7 +422,8 @@
subsubsection{*The Operator @{term is_Forall}, Internalized*}
(* "is_Forall(M,p,Z) == \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" *)
-definition Forall_fm :: "[i,i]=>i"
+definition
+ Forall_fm :: "[i,i]=>i" where
"Forall_fm(x,Z) ==
Exists(And(Inr_fm(succ(x),0), Inr_fm(0,succ(Z))))"
@@ -442,7 +455,8 @@
(* is_and(M,a,b,z) == (number1(M,a) & z=b) |
(~number1(M,a) & empty(M,z)) *)
-definition and_fm :: "[i,i,i]=>i"
+definition
+ and_fm :: "[i,i,i]=>i" where
"and_fm(a,b,z) ==
Or(And(number1_fm(a), Equal(z,b)),
And(Neg(number1_fm(a)),empty_fm(z)))"
@@ -476,7 +490,8 @@
(* is_or(M,a,b,z) == (number1(M,a) & number1(M,z)) |
(~number1(M,a) & z=b) *)
-definition or_fm :: "[i,i,i]=>i"
+definition
+ or_fm :: "[i,i,i]=>i" where
"or_fm(a,b,z) ==
Or(And(number1_fm(a), number1_fm(z)),
And(Neg(number1_fm(a)), Equal(z,b)))"
@@ -510,7 +525,8 @@
(* is_not(M,a,z) == (number1(M,a) & empty(M,z)) |
(~number1(M,a) & number1(M,z)) *)
-definition not_fm :: "[i,i]=>i"
+definition
+ not_fm :: "[i,i]=>i" where
"not_fm(a,z) ==
Or(And(number1_fm(a), empty_fm(z)),
And(Neg(number1_fm(a)), number1_fm(z)))"
@@ -576,8 +592,9 @@
*)
text{*The three arguments of @{term p} are always 2, 1, 0 and z*}
-definition is_recfun_fm :: "[i, i, i, i]=>i"
- "is_recfun_fm(p,r,a,f) ==
+definition
+ is_recfun_fm :: "[i, i, i, i]=>i" where
+ "is_recfun_fm(p,r,a,f) ==
Forall(Iff(Member(0,succ(f)),
Exists(Exists(Exists(
And(p,
@@ -638,8 +655,9 @@
(* is_wfrec :: "[i=>o, i, [i,i,i]=>o, i, i] => o"
"is_wfrec(M,MH,r,a,z) ==
\<exists>f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)" *)
-definition is_wfrec_fm :: "[i, i, i, i]=>i"
- "is_wfrec_fm(p,r,a,z) ==
+definition
+ is_wfrec_fm :: "[i, i, i, i]=>i" where
+ "is_wfrec_fm(p,r,a,z) ==
Exists(And(is_recfun_fm(p, succ(r), succ(a), 0),
Exists(Exists(Exists(Exists(
And(Equal(2,a#+5), And(Equal(1,4), And(Equal(0,z#+5), p)))))))))"
@@ -696,7 +714,8 @@
subsubsection{*Binary Products, Internalized*}
-definition cartprod_fm :: "[i,i,i]=>i"
+definition
+ cartprod_fm :: "[i,i,i]=>i" where
(* "cartprod(M,A,B,z) ==
\<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))" *)
"cartprod_fm(A,B,z) ==
@@ -736,7 +755,8 @@
3 2 1 0
number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) &
cartprod(M,s1,B,B1) & union(M,A0,B1,Z)" *)
-definition sum_fm :: "[i,i,i]=>i"
+definition
+ sum_fm :: "[i,i,i]=>i" where
"sum_fm(A,B,Z) ==
Exists(Exists(Exists(Exists(
And(number1_fm(2),
@@ -771,7 +791,8 @@
subsubsection{*The Operator @{term quasinat}*}
(* "is_quasinat(M,z) == empty(M,z) | (\<exists>m[M]. successor(M,m,z))" *)
-definition quasinat_fm :: "i=>i"
+definition
+ quasinat_fm :: "i=>i" where
"quasinat_fm(z) == Or(empty_fm(z), Exists(succ_fm(0,succ(z))))"
lemma quasinat_type [TC]:
@@ -808,7 +829,8 @@
(\<forall>m[M]. successor(M,m,k) --> is_b(m,z)) &
(is_quasinat(M,k) | empty(M,z))" *)
text{*The formula @{term is_b} has free variables 1 and 0.*}
-definition is_nat_case_fm :: "[i, i, i, i]=>i"
+definition
+ is_nat_case_fm :: "[i, i, i, i]=>i" where
"is_nat_case_fm(a,is_b,k,z) ==
And(Implies(empty_fm(k), Equal(z,a)),
And(Forall(Implies(succ_fm(0,succ(k)),
@@ -863,7 +885,8 @@
"iterates_MH(M,isF,v,n,g,z) ==
is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) & isF(gm,u),
n, z)" *)
-definition iterates_MH_fm :: "[i, i, i, i, i]=>i"
+definition
+ iterates_MH_fm :: "[i, i, i, i, i]=>i" where
"iterates_MH_fm(isF,v,n,g,z) ==
is_nat_case_fm(v,
Exists(And(fun_apply_fm(succ(succ(succ(g))),2,0),
@@ -928,8 +951,9 @@
\<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) & membership(M,sn,msn) &
1 0 is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)"*)
-definition is_iterates_fm :: "[i, i, i, i]=>i"
- "is_iterates_fm(p,v,n,Z) ==
+definition
+ is_iterates_fm :: "[i, i, i, i]=>i" where
+ "is_iterates_fm(p,v,n,Z) ==
Exists(Exists(
And(succ_fm(n#+2,1),
And(Memrel_fm(1,0),
@@ -998,7 +1022,8 @@
(* is_eclose_n(M,A,n,Z) == is_iterates(M, big_union(M), A, n, Z) *)
-definition eclose_n_fm :: "[i,i,i]=>i"
+definition
+ eclose_n_fm :: "[i,i,i]=>i" where
"eclose_n_fm(A,n,Z) == is_iterates_fm(big_union_fm(1,0), A, n, Z)"
lemma eclose_n_fm_type [TC]:
@@ -1034,7 +1059,8 @@
(* mem_eclose(M,A,l) ==
\<exists>n[M]. \<exists>eclosen[M].
finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \<in> eclosen *)
-definition mem_eclose_fm :: "[i,i]=>i"
+definition
+ mem_eclose_fm :: "[i,i]=>i" where
"mem_eclose_fm(x,y) ==
Exists(Exists(
And(finite_ordinal_fm(1),
@@ -1066,7 +1092,8 @@
subsubsection{*The Predicate ``Is @{term "eclose(A)"}''*}
(* is_eclose(M,A,Z) == \<forall>l[M]. l \<in> Z <-> mem_eclose(M,A,l) *)
-definition is_eclose_fm :: "[i,i]=>i"
+definition
+ is_eclose_fm :: "[i,i]=>i" where
"is_eclose_fm(A,Z) ==
Forall(Iff(Member(0,succ(Z)), mem_eclose_fm(succ(A),0)))"
@@ -1095,7 +1122,8 @@
subsubsection{*The List Functor, Internalized*}
-definition list_functor_fm :: "[i,i,i]=>i"
+definition
+ list_functor_fm :: "[i,i,i]=>i" where
(* "is_list_functor(M,A,X,Z) ==
\<exists>n1[M]. \<exists>AX[M].
number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" *)
@@ -1135,7 +1163,8 @@
\<exists>zero[M]. empty(M,zero) &
is_iterates(M, is_list_functor(M,A), zero, n, Z)" *)
-definition list_N_fm :: "[i,i,i]=>i"
+definition
+ list_N_fm :: "[i,i,i]=>i" where
"list_N_fm(A,n,Z) ==
Exists(
And(empty_fm(0),
@@ -1175,7 +1204,8 @@
(* mem_list(M,A,l) ==
\<exists>n[M]. \<exists>listn[M].
finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \<in> listn *)
-definition mem_list_fm :: "[i,i]=>i"
+definition
+ mem_list_fm :: "[i,i]=>i" where
"mem_list_fm(x,y) ==
Exists(Exists(
And(finite_ordinal_fm(1),
@@ -1207,7 +1237,8 @@
subsubsection{*The Predicate ``Is @{term "list(A)"}''*}
(* is_list(M,A,Z) == \<forall>l[M]. l \<in> Z <-> mem_list(M,A,l) *)
-definition is_list_fm :: "[i,i]=>i"
+definition
+ is_list_fm :: "[i,i]=>i" where
"is_list_fm(A,Z) ==
Forall(Iff(Member(0,succ(Z)), mem_list_fm(succ(A),0)))"
@@ -1236,7 +1267,7 @@
subsubsection{*The Formula Functor, Internalized*}
-definition formula_functor_fm :: "[i,i]=>i"
+definition formula_functor_fm :: "[i,i]=>i" where
(* "is_formula_functor(M,X,Z) ==
\<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M].
4 3 2 1 0
@@ -1282,7 +1313,8 @@
(* "is_formula_N(M,n,Z) ==
\<exists>zero[M]. empty(M,zero) &
is_iterates(M, is_formula_functor(M), zero, n, Z)" *)
-definition formula_N_fm :: "[i,i]=>i"
+definition
+ formula_N_fm :: "[i,i]=>i" where
"formula_N_fm(n,Z) ==
Exists(
And(empty_fm(0),
@@ -1322,7 +1354,8 @@
(* mem_formula(M,p) ==
\<exists>n[M]. \<exists>formn[M].
finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \<in> formn *)
-definition mem_formula_fm :: "i=>i"
+definition
+ mem_formula_fm :: "i=>i" where
"mem_formula_fm(x) ==
Exists(Exists(
And(finite_ordinal_fm(1),
@@ -1354,7 +1387,8 @@
subsubsection{*The Predicate ``Is @{term "formula"}''*}
(* is_formula(M,Z) == \<forall>p[M]. p \<in> Z <-> mem_formula(M,p) *)
-definition is_formula_fm :: "i=>i"
+definition
+ is_formula_fm :: "i=>i" where
"is_formula_fm(Z) == Forall(Iff(Member(0,succ(Z)), mem_formula_fm(0)))"
lemma is_formula_type [TC]:
@@ -1392,7 +1426,8 @@
2 1 0
upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) &
is_wfrec(M,MH,mesa,a,z)" *)
-definition is_transrec_fm :: "[i, i, i]=>i"
+definition
+ is_transrec_fm :: "[i, i, i]=>i" where
"is_transrec_fm(p,a,z) ==
Exists(Exists(Exists(
And(upair_fm(a#+3,a#+3,2),
--- a/src/ZF/Constructible/L_axioms.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/L_axioms.thy Fri Nov 17 02:20:03 2006 +0100
@@ -114,21 +114,24 @@
subsection{*Instantiation of the locale @{text reflection}*}
text{*instances of locale constants*}
+
definition
- L_F0 :: "[i=>o,i] => i"
+ L_F0 :: "[i=>o,i] => i" where
"L_F0(P,y) == \<mu> b. (\<exists>z. L(z) \<and> P(<y,z>)) --> (\<exists>z\<in>Lset(b). P(<y,z>))"
- L_FF :: "[i=>o,i] => i"
+definition
+ L_FF :: "[i=>o,i] => i" where
"L_FF(P) == \<lambda>a. \<Union>y\<in>Lset(a). L_F0(P,y)"
- L_ClEx :: "[i=>o,i] => o"
+definition
+ L_ClEx :: "[i=>o,i] => o" where
"L_ClEx(P) == \<lambda>a. Limit(a) \<and> normalize(L_FF(P),a) = a"
text{*We must use the meta-existential quantifier; otherwise the reflection
terms become enormous!*}
definition
- L_Reflects :: "[i=>o,[i,i]=>o] => prop" ("(3REFLECTS/ [_,/ _])")
+ L_Reflects :: "[i=>o,[i,i]=>o] => prop" ("(3REFLECTS/ [_,/ _])") where
"REFLECTS[P,Q] == (??Cl. Closed_Unbounded(Cl) &
(\<forall>a. Cl(a) --> (\<forall>x \<in> Lset(a). P(x) <-> Q(a,x))))"
@@ -266,25 +269,31 @@
subsubsection{*Some numbers to help write de Bruijn indices*}
abbreviation
- digit3 :: i ("3")
- "3 == succ(2)"
- digit4 :: i ("4")
- "4 == succ(3)"
- digit5 :: i ("5")
- "5 == succ(4)"
- digit6 :: i ("6")
- "6 == succ(5)"
- digit7 :: i ("7")
- "7 == succ(6)"
- digit8 :: i ("8")
- "8 == succ(7)"
- digit9 :: i ("9")
- "9 == succ(8)"
+ digit3 :: i ("3") where "3 == succ(2)"
+
+abbreviation
+ digit4 :: i ("4") where "4 == succ(3)"
+
+abbreviation
+ digit5 :: i ("5") where "5 == succ(4)"
+
+abbreviation
+ digit6 :: i ("6") where "6 == succ(5)"
+
+abbreviation
+ digit7 :: i ("7") where "7 == succ(6)"
+
+abbreviation
+ digit8 :: i ("8") where "8 == succ(7)"
+
+abbreviation
+ digit9 :: i ("9") where "9 == succ(8)"
subsubsection{*The Empty Set, Internalized*}
-definition empty_fm :: "i=>i"
+definition
+ empty_fm :: "i=>i" where
"empty_fm(x) == Forall(Neg(Member(0,succ(x))))"
lemma empty_type [TC]:
@@ -322,7 +331,8 @@
subsubsection{*Unordered Pairs, Internalized*}
-definition upair_fm :: "[i,i,i]=>i"
+definition
+ upair_fm :: "[i,i,i]=>i" where
"upair_fm(x,y,z) ==
And(Member(x,z),
And(Member(y,z),
@@ -364,7 +374,8 @@
subsubsection{*Ordered pairs, Internalized*}
-definition pair_fm :: "[i,i,i]=>i"
+definition
+ pair_fm :: "[i,i,i]=>i" where
"pair_fm(x,y,z) ==
Exists(And(upair_fm(succ(x),succ(x),0),
Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0),
@@ -396,7 +407,8 @@
subsubsection{*Binary Unions, Internalized*}
-definition union_fm :: "[i,i,i]=>i"
+definition
+ union_fm :: "[i,i,i]=>i" where
"union_fm(x,y,z) ==
Forall(Iff(Member(0,succ(z)),
Or(Member(0,succ(x)),Member(0,succ(y)))))"
@@ -427,7 +439,8 @@
subsubsection{*Set ``Cons,'' Internalized*}
-definition cons_fm :: "[i,i,i]=>i"
+definition
+ cons_fm :: "[i,i,i]=>i" where
"cons_fm(x,y,z) ==
Exists(And(upair_fm(succ(x),succ(x),0),
union_fm(0,succ(y),succ(z))))"
@@ -459,7 +472,8 @@
subsubsection{*Successor Function, Internalized*}
-definition succ_fm :: "[i,i]=>i"
+definition
+ succ_fm :: "[i,i]=>i" where
"succ_fm(x,y) == cons_fm(x,x,y)"
lemma succ_type [TC]:
@@ -489,7 +503,8 @@
subsubsection{*The Number 1, Internalized*}
(* "number1(M,a) == (\<exists>x[M]. empty(M,x) & successor(M,x,a))" *)
-definition number1_fm :: "i=>i"
+definition
+ number1_fm :: "i=>i" where
"number1_fm(a) == Exists(And(empty_fm(0), succ_fm(0,succ(a))))"
lemma number1_type [TC]:
@@ -518,7 +533,8 @@
subsubsection{*Big Union, Internalized*}
(* "big_union(M,A,z) == \<forall>x[M]. x \<in> z <-> (\<exists>y[M]. y\<in>A & x \<in> y)" *)
-definition big_union_fm :: "[i,i]=>i"
+definition
+ big_union_fm :: "[i,i]=>i" where
"big_union_fm(A,z) ==
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(A))), Member(1,0)))))"
@@ -598,7 +614,8 @@
subsubsection{*Membership Relation, Internalized*}
-definition Memrel_fm :: "[i,i]=>i"
+definition
+ Memrel_fm :: "[i,i]=>i" where
"Memrel_fm(A,r) ==
Forall(Iff(Member(0,succ(r)),
Exists(And(Member(0,succ(succ(A))),
@@ -631,7 +648,8 @@
subsubsection{*Predecessor Set, Internalized*}
-definition pred_set_fm :: "[i,i,i,i]=>i"
+definition
+ pred_set_fm :: "[i,i,i,i]=>i" where
"pred_set_fm(A,x,r,B) ==
Forall(Iff(Member(0,succ(B)),
Exists(And(Member(0,succ(succ(r))),
@@ -669,7 +687,8 @@
(* "is_domain(M,r,z) ==
\<forall>x[M]. (x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w))))" *)
-definition domain_fm :: "[i,i]=>i"
+definition
+ domain_fm :: "[i,i]=>i" where
"domain_fm(r,z) ==
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(r))),
@@ -703,7 +722,8 @@
(* "is_range(M,r,z) ==
\<forall>y[M]. (y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w))))" *)
-definition range_fm :: "[i,i]=>i"
+definition
+ range_fm :: "[i,i]=>i" where
"range_fm(r,z) ==
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(r))),
@@ -738,7 +758,8 @@
(* "is_field(M,r,z) ==
\<exists>dr[M]. is_domain(M,r,dr) &
(\<exists>rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))" *)
-definition field_fm :: "[i,i]=>i"
+definition
+ field_fm :: "[i,i]=>i" where
"field_fm(r,z) ==
Exists(And(domain_fm(succ(r),0),
Exists(And(range_fm(succ(succ(r)),0),
@@ -773,7 +794,8 @@
(* "image(M,r,A,z) ==
\<forall>y[M]. (y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w))))" *)
-definition image_fm :: "[i,i,i]=>i"
+definition
+ image_fm :: "[i,i,i]=>i" where
"image_fm(r,A,z) ==
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(r))),
@@ -808,7 +830,8 @@
(* "pre_image(M,r,A,z) ==
\<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))" *)
-definition pre_image_fm :: "[i,i,i]=>i"
+definition
+ pre_image_fm :: "[i,i,i]=>i" where
"pre_image_fm(r,A,z) ==
Forall(Iff(Member(0,succ(z)),
Exists(And(Member(0,succ(succ(r))),
@@ -844,7 +867,8 @@
(* "fun_apply(M,f,x,y) ==
(\<exists>xs[M]. \<exists>fxs[M].
upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))" *)
-definition fun_apply_fm :: "[i,i,i]=>i"
+definition
+ fun_apply_fm :: "[i,i,i]=>i" where
"fun_apply_fm(f,x,y) ==
Exists(Exists(And(upair_fm(succ(succ(x)), succ(succ(x)), 1),
And(image_fm(succ(succ(f)), 1, 0),
@@ -879,7 +903,8 @@
(* "is_relation(M,r) ==
(\<forall>z[M]. z\<in>r --> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))" *)
-definition relation_fm :: "i=>i"
+definition
+ relation_fm :: "i=>i" where
"relation_fm(r) ==
Forall(Implies(Member(0,succ(r)), Exists(Exists(pair_fm(1,0,2)))))"
@@ -911,7 +936,8 @@
(* "is_function(M,r) ==
\<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
pair(M,x,y,p) --> pair(M,x,y',p') --> p\<in>r --> p'\<in>r --> y=y'" *)
-definition function_fm :: "i=>i"
+definition
+ function_fm :: "i=>i" where
"function_fm(r) ==
Forall(Forall(Forall(Forall(Forall(
Implies(pair_fm(4,3,1),
@@ -948,7 +974,8 @@
is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
(\<forall>u[M]. u\<in>r --> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) --> y\<in>B))" *)
-definition typed_function_fm :: "[i,i,i]=>i"
+definition
+ typed_function_fm :: "[i,i,i]=>i" where
"typed_function_fm(A,B,r) ==
And(function_fm(r),
And(relation_fm(r),
@@ -1007,7 +1034,8 @@
(\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) &
xy \<in> s & yz \<in> r)" *)
-definition composition_fm :: "[i,i,i]=>i"
+definition
+ composition_fm :: "[i,i,i]=>i" where
"composition_fm(r,s,t) ==
Forall(Iff(Member(0,succ(t)),
Exists(Exists(Exists(Exists(Exists(
@@ -1046,8 +1074,9 @@
typed_function(M,A,B,f) &
(\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
pair(M,x,y,p) --> pair(M,x',y,p') --> p\<in>f --> p'\<in>f --> x=x')" *)
-definition injection_fm :: "[i,i,i]=>i"
- "injection_fm(A,B,f) ==
+definition
+ injection_fm :: "[i,i,i]=>i" where
+ "injection_fm(A,B,f) ==
And(typed_function_fm(A,B,f),
Forall(Forall(Forall(Forall(Forall(
Implies(pair_fm(4,2,1),
@@ -1086,8 +1115,9 @@
"surjection(M,A,B,f) ==
typed_function(M,A,B,f) &
(\<forall>y[M]. y\<in>B --> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))" *)
-definition surjection_fm :: "[i,i,i]=>i"
- "surjection_fm(A,B,f) ==
+definition
+ surjection_fm :: "[i,i,i]=>i" where
+ "surjection_fm(A,B,f) ==
And(typed_function_fm(A,B,f),
Forall(Implies(Member(0,succ(B)),
Exists(And(Member(0,succ(succ(A))),
@@ -1122,8 +1152,9 @@
(* bijection :: "[i=>o,i,i,i] => o"
"bijection(M,A,B,f) == injection(M,A,B,f) & surjection(M,A,B,f)" *)
-definition bijection_fm :: "[i,i,i]=>i"
- "bijection_fm(A,B,f) == And(injection_fm(A,B,f), surjection_fm(A,B,f))"
+definition
+ bijection_fm :: "[i,i,i]=>i" where
+ "bijection_fm(A,B,f) == And(injection_fm(A,B,f), surjection_fm(A,B,f))"
lemma bijection_type [TC]:
"[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> bijection_fm(x,y,z) \<in> formula"
@@ -1154,7 +1185,8 @@
(* "restriction(M,r,A,z) ==
\<forall>x[M]. x \<in> z <-> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))" *)
-definition restriction_fm :: "[i,i,i]=>i"
+definition
+ restriction_fm :: "[i,i,i]=>i" where
"restriction_fm(r,A,z) ==
Forall(Iff(Member(0,succ(z)),
And(Member(0,succ(r)),
@@ -1195,7 +1227,8 @@
pair(M,fx,fy,q) --> (p\<in>r <-> q\<in>s))))"
*)
-definition order_isomorphism_fm :: "[i,i,i,i,i]=>i"
+definition
+ order_isomorphism_fm :: "[i,i,i,i,i]=>i" where
"order_isomorphism_fm(A,r,B,s,f) ==
And(bijection_fm(A,B,f),
Forall(Implies(Member(0,succ(A)),
@@ -1242,7 +1275,8 @@
ordinal(M,a) & ~ empty(M,a) &
(\<forall>x[M]. x\<in>a --> (\<exists>y[M]. y\<in>a & successor(M,x,y)))" *)
-definition limit_ordinal_fm :: "i=>i"
+definition
+ limit_ordinal_fm :: "i=>i" where
"limit_ordinal_fm(x) ==
And(ordinal_fm(x),
And(Neg(empty_fm(x)),
@@ -1278,7 +1312,8 @@
(* "finite_ordinal(M,a) ==
ordinal(M,a) & ~ limit_ordinal(M,a) &
(\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x))" *)
-definition finite_ordinal_fm :: "i=>i"
+definition
+ finite_ordinal_fm :: "i=>i" where
"finite_ordinal_fm(x) ==
And(ordinal_fm(x),
And(Neg(limit_ordinal_fm(x)),
@@ -1311,7 +1346,8 @@
subsubsection{*Omega: The Set of Natural Numbers*}
(* omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x)) *)
-definition omega_fm :: "i=>i"
+definition
+ omega_fm :: "i=>i" where
"omega_fm(x) ==
And(limit_ordinal_fm(x),
Forall(Implies(Member(0,succ(x)),
--- a/src/ZF/Constructible/MetaExists.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/MetaExists.thy Fri Nov 17 02:20:03 2006 +0100
@@ -11,7 +11,7 @@
quantify over classes. Yields a proposition rather than a FOL formula.*}
definition
- ex :: "(('a::{}) => prop) => prop" (binder "?? " 0)
+ ex :: "(('a::{}) => prop) => prop" (binder "?? " 0) where
"ex(P) == (!!Q. (!!x. PROP P(x) ==> PROP Q) ==> PROP Q)"
notation (xsymbols)
--- a/src/ZF/Constructible/Normal.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/Normal.thy Fri Nov 17 02:20:03 2006 +0100
@@ -19,13 +19,15 @@
subsection {*Closed and Unbounded (c.u.) Classes of Ordinals*}
definition
- Closed :: "(i=>o) => o"
+ Closed :: "(i=>o) => o" where
"Closed(P) == \<forall>I. I \<noteq> 0 --> (\<forall>i\<in>I. Ord(i) \<and> P(i)) --> P(\<Union>(I))"
- Unbounded :: "(i=>o) => o"
+definition
+ Unbounded :: "(i=>o) => o" where
"Unbounded(P) == \<forall>i. Ord(i) --> (\<exists>j. i<j \<and> P(j))"
- Closed_Unbounded :: "(i=>o) => o"
+definition
+ Closed_Unbounded :: "(i=>o) => o" where
"Closed_Unbounded(P) == Closed(P) \<and> Unbounded(P)"
@@ -201,16 +203,19 @@
subsection {*Normal Functions*}
definition
- mono_le_subset :: "(i=>i) => o"
+ mono_le_subset :: "(i=>i) => o" where
"mono_le_subset(M) == \<forall>i j. i\<le>j --> M(i) \<subseteq> M(j)"
- mono_Ord :: "(i=>i) => o"
+definition
+ mono_Ord :: "(i=>i) => o" where
"mono_Ord(F) == \<forall>i j. i<j --> F(i) < F(j)"
- cont_Ord :: "(i=>i) => o"
+definition
+ cont_Ord :: "(i=>i) => o" where
"cont_Ord(F) == \<forall>l. Limit(l) --> F(l) = (\<Union>i<l. F(i))"
- Normal :: "(i=>i) => o"
+definition
+ Normal :: "(i=>i) => o" where
"Normal(F) == mono_Ord(F) \<and> cont_Ord(F)"
@@ -373,7 +378,7 @@
normal function, by @{thm [source] Normal_imp_fp_Unbounded}.
*}
definition
- normalize :: "[i=>i, i] => i"
+ normalize :: "[i=>i, i] => i" where
"normalize(F,a) == transrec2(a, F(0), \<lambda>x r. F(succ(x)) Un succ(r))"
@@ -425,7 +430,7 @@
numbers.*}
definition
- Aleph :: "i => i"
+ Aleph :: "i => i" where
"Aleph(a) == transrec2(a, nat, \<lambda>x r. csucc(r))"
notation (xsymbols)
--- a/src/ZF/Constructible/Rank.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/Rank.thy Fri Nov 17 02:20:03 2006 +0100
@@ -136,22 +136,22 @@
done
-definition
-
- obase :: "[i=>o,i,i] => i"
+definition
+ obase :: "[i=>o,i,i] => i" where
--{*the domain of @{text om}, eventually shown to equal @{text A}*}
"obase(M,A,r) == {a\<in>A. \<exists>x[M]. \<exists>g[M]. Ord(x) &
g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x))}"
- omap :: "[i=>o,i,i,i] => o"
+definition
+ omap :: "[i=>o,i,i,i] => o" where
--{*the function that maps wosets to order types*}
"omap(M,A,r,f) ==
\<forall>z[M].
z \<in> f <-> (\<exists>a\<in>A. \<exists>x[M]. \<exists>g[M]. z = <a,x> & Ord(x) &
g \<in> ord_iso(Order.pred(A,a,r),r,x,Memrel(x)))"
-
- otype :: "[i=>o,i,i,i] => o" --{*the order types themselves*}
+definition
+ otype :: "[i=>o,i,i,i] => o" where --{*the order types themselves*}
"otype(M,A,r,i) == \<exists>f[M]. omap(M,A,r,f) & is_range(M,f,i)"
@@ -414,12 +414,12 @@
subsubsection{*Ordinal Addition*}
(*FIXME: update to use new techniques!!*)
-definition
(*This expresses ordinal addition in the language of ZF. It also
provides an abbreviation that can be used in the instance of strong
replacement below. Here j is used to define the relation, namely
Memrel(succ(j)), while x determines the domain of f.*)
- is_oadd_fun :: "[i=>o,i,i,i,i] => o"
+definition
+ is_oadd_fun :: "[i=>o,i,i,i,i] => o" where
"is_oadd_fun(M,i,j,x,f) ==
(\<forall>sj msj. M(sj) --> M(msj) -->
successor(M,j,sj) --> membership(M,sj,msj) -->
@@ -427,7 +427,8 @@
%x g y. \<exists>gx[M]. image(M,g,x,gx) & union(M,i,gx,y),
msj, x, f))"
- is_oadd :: "[i=>o,i,i,i] => o"
+definition
+ is_oadd :: "[i=>o,i,i,i] => o" where
"is_oadd(M,i,j,k) ==
(~ ordinal(M,i) & ~ ordinal(M,j) & k=0) |
(~ ordinal(M,i) & ordinal(M,j) & k=j) |
@@ -437,21 +438,24 @@
successor(M,j,sj) & is_oadd_fun(M,i,sj,sj,f) &
fun_apply(M,f,j,fj) & fj = k))"
+definition
(*NEEDS RELATIVIZATION*)
- omult_eqns :: "[i,i,i,i] => o"
+ omult_eqns :: "[i,i,i,i] => o" where
"omult_eqns(i,x,g,z) ==
Ord(x) &
(x=0 --> z=0) &
(\<forall>j. x = succ(j) --> z = g`j ++ i) &
(Limit(x) --> z = \<Union>(g``x))"
- is_omult_fun :: "[i=>o,i,i,i] => o"
+definition
+ is_omult_fun :: "[i=>o,i,i,i] => o" where
"is_omult_fun(M,i,j,f) ==
(\<exists>df. M(df) & is_function(M,f) &
is_domain(M,f,df) & subset(M, j, df)) &
(\<forall>x\<in>j. omult_eqns(i,x,f,f`x))"
- is_omult :: "[i=>o,i,i,i] => o"
+definition
+ is_omult :: "[i=>o,i,i,i] => o" where
"is_omult(M,i,j,k) ==
\<exists>f fj sj. M(f) & M(fj) & M(sj) &
successor(M,j,sj) & is_omult_fun(M,i,sj,f) &
@@ -726,7 +730,7 @@
text{*This function, defined using replacement, is a rank function for
well-founded relations within the class M.*}
definition
- wellfoundedrank :: "[i=>o,i,i] => i"
+ wellfoundedrank :: "[i=>o,i,i] => i" where
"wellfoundedrank(M,r,A) ==
{p. x\<in>A, \<exists>y[M]. \<exists>f[M].
p = <x,y> & is_recfun(r^+, x, %x f. range(f), f) &
--- a/src/ZF/Constructible/Rec_Separation.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/Rec_Separation.thy Fri Nov 17 02:20:03 2006 +0100
@@ -30,7 +30,8 @@
(\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M].
fun_apply(M,f,j,fj) & successor(M,j,sj) &
fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"*)
-definition rtran_closure_mem_fm :: "[i,i,i]=>i"
+definition
+ rtran_closure_mem_fm :: "[i,i,i]=>i" where
"rtran_closure_mem_fm(A,r,p) ==
Exists(Exists(Exists(
And(omega_fm(2),
@@ -87,8 +88,9 @@
(* "rtran_closure(M,r,s) ==
\<forall>A[M]. is_field(M,r,A) -->
(\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))" *)
-definition rtran_closure_fm :: "[i,i]=>i"
- "rtran_closure_fm(r,s) ==
+definition
+ rtran_closure_fm :: "[i,i]=>i" where
+ "rtran_closure_fm(r,s) ==
Forall(Implies(field_fm(succ(r),0),
Forall(Iff(Member(0,succ(succ(s))),
rtran_closure_mem_fm(1,succ(succ(r)),0)))))"
@@ -121,8 +123,9 @@
(* "tran_closure(M,r,t) ==
\<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *)
-definition tran_closure_fm :: "[i,i]=>i"
- "tran_closure_fm(r,s) ==
+definition
+ tran_closure_fm :: "[i,i]=>i" where
+ "tran_closure_fm(r,s) ==
Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))"
lemma tran_closure_type [TC]:
@@ -293,7 +296,8 @@
(* "is_nth(M,n,l,Z) ==
\<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)" *)
-definition nth_fm :: "[i,i,i]=>i"
+definition
+ nth_fm :: "[i,i,i]=>i" where
"nth_fm(n,l,Z) ==
Exists(And(is_iterates_fm(tl_fm(1,0), succ(l), succ(n), 0),
hd_fm(0,succ(Z))))"
--- a/src/ZF/Constructible/Relative.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/Relative.thy Fri Nov 17 02:20:03 2006 +0100
@@ -10,95 +10,120 @@
subsection{* Relativized versions of standard set-theoretic concepts *}
definition
- empty :: "[i=>o,i] => o"
+ empty :: "[i=>o,i] => o" where
"empty(M,z) == \<forall>x[M]. x \<notin> z"
- subset :: "[i=>o,i,i] => o"
+definition
+ subset :: "[i=>o,i,i] => o" where
"subset(M,A,B) == \<forall>x[M]. x\<in>A --> x \<in> B"
- upair :: "[i=>o,i,i,i] => o"
+definition
+ upair :: "[i=>o,i,i,i] => o" where
"upair(M,a,b,z) == a \<in> z & b \<in> z & (\<forall>x[M]. x\<in>z --> x = a | x = b)"
- pair :: "[i=>o,i,i,i] => o"
+definition
+ pair :: "[i=>o,i,i,i] => o" where
"pair(M,a,b,z) == \<exists>x[M]. upair(M,a,a,x) &
- (\<exists>y[M]. upair(M,a,b,y) & upair(M,x,y,z))"
+ (\<exists>y[M]. upair(M,a,b,y) & upair(M,x,y,z))"
- union :: "[i=>o,i,i,i] => o"
+definition
+ union :: "[i=>o,i,i,i] => o" where
"union(M,a,b,z) == \<forall>x[M]. x \<in> z <-> x \<in> a | x \<in> b"
- is_cons :: "[i=>o,i,i,i] => o"
+definition
+ is_cons :: "[i=>o,i,i,i] => o" where
"is_cons(M,a,b,z) == \<exists>x[M]. upair(M,a,a,x) & union(M,x,b,z)"
- successor :: "[i=>o,i,i] => o"
+definition
+ successor :: "[i=>o,i,i] => o" where
"successor(M,a,z) == is_cons(M,a,a,z)"
- number1 :: "[i=>o,i] => o"
+definition
+ number1 :: "[i=>o,i] => o" where
"number1(M,a) == \<exists>x[M]. empty(M,x) & successor(M,x,a)"
- number2 :: "[i=>o,i] => o"
+definition
+ number2 :: "[i=>o,i] => o" where
"number2(M,a) == \<exists>x[M]. number1(M,x) & successor(M,x,a)"
- number3 :: "[i=>o,i] => o"
+definition
+ number3 :: "[i=>o,i] => o" where
"number3(M,a) == \<exists>x[M]. number2(M,x) & successor(M,x,a)"
- powerset :: "[i=>o,i,i] => o"
+definition
+ powerset :: "[i=>o,i,i] => o" where
"powerset(M,A,z) == \<forall>x[M]. x \<in> z <-> subset(M,x,A)"
- is_Collect :: "[i=>o,i,i=>o,i] => o"
+definition
+ is_Collect :: "[i=>o,i,i=>o,i] => o" where
"is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z <-> x \<in> A & P(x)"
- is_Replace :: "[i=>o,i,[i,i]=>o,i] => o"
+definition
+ is_Replace :: "[i=>o,i,[i,i]=>o,i] => o" where
"is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & P(x,u))"
- inter :: "[i=>o,i,i,i] => o"
+definition
+ inter :: "[i=>o,i,i,i] => o" where
"inter(M,a,b,z) == \<forall>x[M]. x \<in> z <-> x \<in> a & x \<in> b"
- setdiff :: "[i=>o,i,i,i] => o"
+definition
+ setdiff :: "[i=>o,i,i,i] => o" where
"setdiff(M,a,b,z) == \<forall>x[M]. x \<in> z <-> x \<in> a & x \<notin> b"
- big_union :: "[i=>o,i,i] => o"
+definition
+ big_union :: "[i=>o,i,i] => o" where
"big_union(M,A,z) == \<forall>x[M]. x \<in> z <-> (\<exists>y[M]. y\<in>A & x \<in> y)"
- big_inter :: "[i=>o,i,i] => o"
+definition
+ big_inter :: "[i=>o,i,i] => o" where
"big_inter(M,A,z) ==
(A=0 --> z=0) &
(A\<noteq>0 --> (\<forall>x[M]. x \<in> z <-> (\<forall>y[M]. y\<in>A --> x \<in> y)))"
- cartprod :: "[i=>o,i,i,i] => o"
+definition
+ cartprod :: "[i=>o,i,i,i] => o" where
"cartprod(M,A,B,z) ==
\<forall>u[M]. u \<in> z <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))"
- is_sum :: "[i=>o,i,i,i] => o"
+definition
+ is_sum :: "[i=>o,i,i,i] => o" where
"is_sum(M,A,B,Z) ==
\<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) &
cartprod(M,s1,B,B1) & union(M,A0,B1,Z)"
- is_Inl :: "[i=>o,i,i] => o"
+definition
+ is_Inl :: "[i=>o,i,i] => o" where
"is_Inl(M,a,z) == \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z)"
- is_Inr :: "[i=>o,i,i] => o"
+definition
+ is_Inr :: "[i=>o,i,i] => o" where
"is_Inr(M,a,z) == \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z)"
- is_converse :: "[i=>o,i,i] => o"
+definition
+ is_converse :: "[i=>o,i,i] => o" where
"is_converse(M,r,z) ==
\<forall>x[M]. x \<in> z <->
(\<exists>w[M]. w\<in>r & (\<exists>u[M]. \<exists>v[M]. pair(M,u,v,w) & pair(M,v,u,x)))"
- pre_image :: "[i=>o,i,i,i] => o"
+definition
+ pre_image :: "[i=>o,i,i,i] => o" where
"pre_image(M,r,A,z) ==
\<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))"
- is_domain :: "[i=>o,i,i] => o"
+definition
+ is_domain :: "[i=>o,i,i] => o" where
"is_domain(M,r,z) ==
\<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w)))"
- image :: "[i=>o,i,i,i] => o"
+definition
+ image :: "[i=>o,i,i,i] => o" where
"image(M,r,A,z) ==
\<forall>y[M]. y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w)))"
- is_range :: "[i=>o,i,i] => o"
+definition
+ is_range :: "[i=>o,i,i] => o" where
--{*the cleaner
@{term "\<exists>r'[M]. is_converse(M,r,r') & is_domain(M,r',z)"}
unfortunately needs an instance of separation in order to prove
@@ -106,121 +131,147 @@
"is_range(M,r,z) ==
\<forall>y[M]. y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w)))"
- is_field :: "[i=>o,i,i] => o"
+definition
+ is_field :: "[i=>o,i,i] => o" where
"is_field(M,r,z) ==
\<exists>dr[M]. \<exists>rr[M]. is_domain(M,r,dr) & is_range(M,r,rr) &
union(M,dr,rr,z)"
- is_relation :: "[i=>o,i] => o"
+definition
+ is_relation :: "[i=>o,i] => o" where
"is_relation(M,r) ==
(\<forall>z[M]. z\<in>r --> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))"
- is_function :: "[i=>o,i] => o"
+definition
+ is_function :: "[i=>o,i] => o" where
"is_function(M,r) ==
\<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
pair(M,x,y,p) --> pair(M,x,y',p') --> p\<in>r --> p'\<in>r --> y=y'"
- fun_apply :: "[i=>o,i,i,i] => o"
+definition
+ fun_apply :: "[i=>o,i,i,i] => o" where
"fun_apply(M,f,x,y) ==
(\<exists>xs[M]. \<exists>fxs[M].
upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))"
- typed_function :: "[i=>o,i,i,i] => o"
+definition
+ typed_function :: "[i=>o,i,i,i] => o" where
"typed_function(M,A,B,r) ==
is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
(\<forall>u[M]. u\<in>r --> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) --> y\<in>B))"
- is_funspace :: "[i=>o,i,i,i] => o"
+definition
+ is_funspace :: "[i=>o,i,i,i] => o" where
"is_funspace(M,A,B,F) ==
\<forall>f[M]. f \<in> F <-> typed_function(M,A,B,f)"
- composition :: "[i=>o,i,i,i] => o"
+definition
+ composition :: "[i=>o,i,i,i] => o" where
"composition(M,r,s,t) ==
\<forall>p[M]. p \<in> t <->
(\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) &
xy \<in> s & yz \<in> r)"
- injection :: "[i=>o,i,i,i] => o"
+definition
+ injection :: "[i=>o,i,i,i] => o" where
"injection(M,A,B,f) ==
typed_function(M,A,B,f) &
(\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
pair(M,x,y,p) --> pair(M,x',y,p') --> p\<in>f --> p'\<in>f --> x=x')"
- surjection :: "[i=>o,i,i,i] => o"
+definition
+ surjection :: "[i=>o,i,i,i] => o" where
"surjection(M,A,B,f) ==
typed_function(M,A,B,f) &
(\<forall>y[M]. y\<in>B --> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))"
- bijection :: "[i=>o,i,i,i] => o"
+definition
+ bijection :: "[i=>o,i,i,i] => o" where
"bijection(M,A,B,f) == injection(M,A,B,f) & surjection(M,A,B,f)"
- restriction :: "[i=>o,i,i,i] => o"
+definition
+ restriction :: "[i=>o,i,i,i] => o" where
"restriction(M,r,A,z) ==
\<forall>x[M]. x \<in> z <-> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))"
- transitive_set :: "[i=>o,i] => o"
+definition
+ transitive_set :: "[i=>o,i] => o" where
"transitive_set(M,a) == \<forall>x[M]. x\<in>a --> subset(M,x,a)"
- ordinal :: "[i=>o,i] => o"
+definition
+ ordinal :: "[i=>o,i] => o" where
--{*an ordinal is a transitive set of transitive sets*}
"ordinal(M,a) == transitive_set(M,a) & (\<forall>x[M]. x\<in>a --> transitive_set(M,x))"
- limit_ordinal :: "[i=>o,i] => o"
+definition
+ limit_ordinal :: "[i=>o,i] => o" where
--{*a limit ordinal is a non-empty, successor-closed ordinal*}
"limit_ordinal(M,a) ==
ordinal(M,a) & ~ empty(M,a) &
(\<forall>x[M]. x\<in>a --> (\<exists>y[M]. y\<in>a & successor(M,x,y)))"
- successor_ordinal :: "[i=>o,i] => o"
+definition
+ successor_ordinal :: "[i=>o,i] => o" where
--{*a successor ordinal is any ordinal that is neither empty nor limit*}
"successor_ordinal(M,a) ==
ordinal(M,a) & ~ empty(M,a) & ~ limit_ordinal(M,a)"
- finite_ordinal :: "[i=>o,i] => o"
+definition
+ finite_ordinal :: "[i=>o,i] => o" where
--{*an ordinal is finite if neither it nor any of its elements are limit*}
"finite_ordinal(M,a) ==
ordinal(M,a) & ~ limit_ordinal(M,a) &
(\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x))"
- omega :: "[i=>o,i] => o"
+definition
+ omega :: "[i=>o,i] => o" where
--{*omega is a limit ordinal none of whose elements are limit*}
"omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x))"
- is_quasinat :: "[i=>o,i] => o"
+definition
+ is_quasinat :: "[i=>o,i] => o" where
"is_quasinat(M,z) == empty(M,z) | (\<exists>m[M]. successor(M,m,z))"
- is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o"
+definition
+ is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" where
"is_nat_case(M, a, is_b, k, z) ==
(empty(M,k) --> z=a) &
(\<forall>m[M]. successor(M,m,k) --> is_b(m,z)) &
(is_quasinat(M,k) | empty(M,z))"
- relation1 :: "[i=>o, [i,i]=>o, i=>i] => o"
+definition
+ relation1 :: "[i=>o, [i,i]=>o, i=>i] => o" where
"relation1(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. is_f(x,y) <-> y = f(x)"
- Relation1 :: "[i=>o, i, [i,i]=>o, i=>i] => o"
+definition
+ Relation1 :: "[i=>o, i, [i,i]=>o, i=>i] => o" where
--{*as above, but typed*}
"Relation1(M,A,is_f,f) ==
\<forall>x[M]. \<forall>y[M]. x\<in>A --> is_f(x,y) <-> y = f(x)"
- relation2 :: "[i=>o, [i,i,i]=>o, [i,i]=>i] => o"
+definition
+ relation2 :: "[i=>o, [i,i,i]=>o, [i,i]=>i] => o" where
"relation2(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. is_f(x,y,z) <-> z = f(x,y)"
- Relation2 :: "[i=>o, i, i, [i,i,i]=>o, [i,i]=>i] => o"
+definition
+ Relation2 :: "[i=>o, i, i, [i,i,i]=>o, [i,i]=>i] => o" where
"Relation2(M,A,B,is_f,f) ==
\<forall>x[M]. \<forall>y[M]. \<forall>z[M]. x\<in>A --> y\<in>B --> is_f(x,y,z) <-> z = f(x,y)"
- relation3 :: "[i=>o, [i,i,i,i]=>o, [i,i,i]=>i] => o"
+definition
+ relation3 :: "[i=>o, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
"relation3(M,is_f,f) ==
\<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M]. is_f(x,y,z,u) <-> u = f(x,y,z)"
- Relation3 :: "[i=>o, i, i, i, [i,i,i,i]=>o, [i,i,i]=>i] => o"
+definition
+ Relation3 :: "[i=>o, i, i, i, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
"Relation3(M,A,B,C,is_f,f) ==
\<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M].
x\<in>A --> y\<in>B --> z\<in>C --> is_f(x,y,z,u) <-> u = f(x,y,z)"
- relation4 :: "[i=>o, [i,i,i,i,i]=>o, [i,i,i,i]=>i] => o"
+definition
+ relation4 :: "[i=>o, [i,i,i,i,i]=>o, [i,i,i,i]=>i] => o" where
"relation4(M,is_f,f) ==
\<forall>u[M]. \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>a[M]. is_f(u,x,y,z,a) <-> a = f(u,x,y,z)"
@@ -236,13 +287,14 @@
subsection {*The relativized ZF axioms*}
+
definition
-
- extensionality :: "(i=>o) => o"
+ extensionality :: "(i=>o) => o" where
"extensionality(M) ==
\<forall>x[M]. \<forall>y[M]. (\<forall>z[M]. z \<in> x <-> z \<in> y) --> x=y"
- separation :: "[i=>o, i=>o] => o"
+definition
+ separation :: "[i=>o, i=>o] => o" where
--{*The formula @{text P} should only involve parameters
belonging to @{text M} and all its quantifiers must be relativized
to @{text M}. We do not have separation as a scheme; every instance
@@ -250,30 +302,37 @@
"separation(M,P) ==
\<forall>z[M]. \<exists>y[M]. \<forall>x[M]. x \<in> y <-> x \<in> z & P(x)"
- upair_ax :: "(i=>o) => o"
+definition
+ upair_ax :: "(i=>o) => o" where
"upair_ax(M) == \<forall>x[M]. \<forall>y[M]. \<exists>z[M]. upair(M,x,y,z)"
- Union_ax :: "(i=>o) => o"
+definition
+ Union_ax :: "(i=>o) => o" where
"Union_ax(M) == \<forall>x[M]. \<exists>z[M]. big_union(M,x,z)"
- power_ax :: "(i=>o) => o"
+definition
+ power_ax :: "(i=>o) => o" where
"power_ax(M) == \<forall>x[M]. \<exists>z[M]. powerset(M,x,z)"
- univalent :: "[i=>o, i, [i,i]=>o] => o"
+definition
+ univalent :: "[i=>o, i, [i,i]=>o] => o" where
"univalent(M,A,P) ==
\<forall>x[M]. x\<in>A --> (\<forall>y[M]. \<forall>z[M]. P(x,y) & P(x,z) --> y=z)"
- replacement :: "[i=>o, [i,i]=>o] => o"
+definition
+ replacement :: "[i=>o, [i,i]=>o] => o" where
"replacement(M,P) ==
\<forall>A[M]. univalent(M,A,P) -->
(\<exists>Y[M]. \<forall>b[M]. (\<exists>x[M]. x\<in>A & P(x,b)) --> b \<in> Y)"
- strong_replacement :: "[i=>o, [i,i]=>o] => o"
+definition
+ strong_replacement :: "[i=>o, [i,i]=>o] => o" where
"strong_replacement(M,P) ==
\<forall>A[M]. univalent(M,A,P) -->
(\<exists>Y[M]. \<forall>b[M]. b \<in> Y <-> (\<exists>x[M]. x\<in>A & P(x,b)))"
- foundation_ax :: "(i=>o) => o"
+definition
+ foundation_ax :: "(i=>o) => o" where
"foundation_ax(M) ==
\<forall>x[M]. (\<exists>y[M]. y\<in>x) --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & z \<in> y))"
@@ -441,9 +500,9 @@
text{*More constants, for order types*}
+
definition
-
- order_isomorphism :: "[i=>o,i,i,i,i,i] => o"
+ order_isomorphism :: "[i=>o,i,i,i,i,i] => o" where
"order_isomorphism(M,A,r,B,s,f) ==
bijection(M,A,B,f) &
(\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A -->
@@ -451,11 +510,13 @@
pair(M,x,y,p) --> fun_apply(M,f,x,fx) --> fun_apply(M,f,y,fy) -->
pair(M,fx,fy,q) --> (p\<in>r <-> q\<in>s))))"
- pred_set :: "[i=>o,i,i,i,i] => o"
+definition
+ pred_set :: "[i=>o,i,i,i,i] => o" where
"pred_set(M,A,x,r,B) ==
\<forall>y[M]. y \<in> B <-> (\<exists>p[M]. p\<in>r & y \<in> A & pair(M,y,x,p))"
- membership :: "[i=>o,i,i] => o" --{*membership relation*}
+definition
+ membership :: "[i=>o,i,i] => o" where --{*membership relation*}
"membership(M,A,r) ==
\<forall>p[M]. p \<in> r <-> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>A & x\<in>y & pair(M,x,y,p)))"
@@ -713,7 +774,7 @@
subsubsection {*Absoluteness for @{term Lambda}*}
definition
- is_lambda :: "[i=>o, i, [i,i]=>o, i] => o"
+ is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" where
"is_lambda(M, A, is_b, z) ==
\<forall>p[M]. p \<in> z <->
(\<exists>u[M]. \<exists>v[M]. u\<in>A & pair(M,u,v,p) & is_b(u,v))"
@@ -1313,18 +1374,21 @@
subsection{*Relativization and Absoluteness for Boolean Operators*}
definition
- is_bool_of_o :: "[i=>o, o, i] => o"
+ is_bool_of_o :: "[i=>o, o, i] => o" where
"is_bool_of_o(M,P,z) == (P & number1(M,z)) | (~P & empty(M,z))"
- is_not :: "[i=>o, i, i] => o"
+definition
+ is_not :: "[i=>o, i, i] => o" where
"is_not(M,a,z) == (number1(M,a) & empty(M,z)) |
(~number1(M,a) & number1(M,z))"
- is_and :: "[i=>o, i, i, i] => o"
+definition
+ is_and :: "[i=>o, i, i, i] => o" where
"is_and(M,a,b,z) == (number1(M,a) & z=b) |
(~number1(M,a) & empty(M,z))"
- is_or :: "[i=>o, i, i, i] => o"
+definition
+ is_or :: "[i=>o, i, i, i] => o" where
"is_or(M,a,b,z) == (number1(M,a) & number1(M,z)) |
(~number1(M,a) & z=b)"
@@ -1366,12 +1430,12 @@
subsection{*Relativization and Absoluteness for List Operators*}
definition
-
- is_Nil :: "[i=>o, i] => o"
+ is_Nil :: "[i=>o, i] => o" where
--{* because @{term "[] \<equiv> Inl(0)"}*}
"is_Nil(M,xs) == \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs)"
- is_Cons :: "[i=>o,i,i,i] => o"
+definition
+ is_Cons :: "[i=>o,i,i,i] => o" where
--{* because @{term "Cons(a, l) \<equiv> Inr(\<langle>a,l\<rangle>)"}*}
"is_Cons(M,a,l,Z) == \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)"
@@ -1391,34 +1455,39 @@
definition
-
- quasilist :: "i => o"
+ quasilist :: "i => o" where
"quasilist(xs) == xs=Nil | (\<exists>x l. xs = Cons(x,l))"
- is_quasilist :: "[i=>o,i] => o"
+definition
+ is_quasilist :: "[i=>o,i] => o" where
"is_quasilist(M,z) == is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))"
- list_case' :: "[i, [i,i]=>i, i] => i"
+definition
+ list_case' :: "[i, [i,i]=>i, i] => i" where
--{*A version of @{term list_case} that's always defined.*}
"list_case'(a,b,xs) ==
if quasilist(xs) then list_case(a,b,xs) else 0"
- is_list_case :: "[i=>o, i, [i,i,i]=>o, i, i] => o"
+definition
+ is_list_case :: "[i=>o, i, [i,i,i]=>o, i, i] => o" where
--{*Returns 0 for non-lists*}
"is_list_case(M, a, is_b, xs, z) ==
(is_Nil(M,xs) --> z=a) &
(\<forall>x[M]. \<forall>l[M]. is_Cons(M,x,l,xs) --> is_b(x,l,z)) &
(is_quasilist(M,xs) | empty(M,z))"
- hd' :: "i => i"
+definition
+ hd' :: "i => i" where
--{*A version of @{term hd} that's always defined.*}
"hd'(xs) == if quasilist(xs) then hd(xs) else 0"
- tl' :: "i => i"
+definition
+ tl' :: "i => i" where
--{*A version of @{term tl} that's always defined.*}
"tl'(xs) == if quasilist(xs) then tl(xs) else 0"
- is_hd :: "[i=>o,i,i] => o"
+definition
+ is_hd :: "[i=>o,i,i] => o" where
--{* @{term "hd([]) = 0"} no constraints if not a list.
Avoiding implication prevents the simplifier's looping.*}
"is_hd(M,xs,H) ==
@@ -1426,7 +1495,8 @@
(\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | H=x) &
(is_quasilist(M,xs) | empty(M,H))"
- is_tl :: "[i=>o,i,i] => o"
+definition
+ is_tl :: "[i=>o,i,i] => o" where
--{* @{term "tl([]) = []"}; see comments about @{term is_hd}*}
"is_tl(M,xs,T) ==
(is_Nil(M,xs) --> T=xs) &
--- a/src/ZF/Constructible/Satisfies_absolute.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/Satisfies_absolute.thy Fri Nov 17 02:20:03 2006 +0100
@@ -17,7 +17,8 @@
2 1 0
is_formula_N(M,n,formula_n) & p \<notin> formula_n &
successor(M,n,sn) & is_formula_N(M,sn,formula_sn) & p \<in> formula_sn" *)
-definition depth_fm :: "[i,i]=>i"
+definition
+ depth_fm :: "[i,i]=>i" where
"depth_fm(p,n) ==
Exists(Exists(Exists(
And(formula_N_fm(n#+3,1),
@@ -66,8 +67,9 @@
is_Nand(M,x,y,v) --> is_c(x,y,z)) &
(\<forall>x[M]. x\<in>formula --> is_Forall(M,x,v) --> is_d(x,z))" *)
-definition formula_case_fm :: "[i, i, i, i, i, i]=>i"
- "formula_case_fm(is_a, is_b, is_c, is_d, v, z) ==
+definition
+ formula_case_fm :: "[i, i, i, i, i, i]=>i" where
+ "formula_case_fm(is_a, is_b, is_c, is_d, v, z) ==
And(Forall(Forall(Implies(finite_ordinal_fm(1),
Implies(finite_ordinal_fm(0),
Implies(Member_fm(1,0,v#+2),
@@ -174,9 +176,9 @@
subsection {*Absoluteness for the Function @{term satisfies}*}
definition
- is_depth_apply :: "[i=>o,i,i,i] => o"
+ is_depth_apply :: "[i=>o,i,i,i] => o" where
--{*Merely a useful abbreviation for the sequel.*}
- "is_depth_apply(M,h,p,z) ==
+ "is_depth_apply(M,h,p,z) ==
\<exists>dp[M]. \<exists>sdp[M]. \<exists>hsdp[M].
finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z)"
@@ -194,11 +196,12 @@
text{*These constants let us instantiate the parameters @{term a}, @{term b},
@{term c}, @{term d}, etc., of the locale @{text Formula_Rec}.*}
definition
- satisfies_a :: "[i,i,i]=>i"
+ satisfies_a :: "[i,i,i]=>i" where
"satisfies_a(A) ==
\<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) \<in> nth(y,env))"
- satisfies_is_a :: "[i=>o,i,i,i,i]=>o"
+definition
+ satisfies_is_a :: "[i=>o,i,i,i,i]=>o" where
"satisfies_is_a(M,A) ==
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) -->
is_lambda(M, lA,
@@ -207,24 +210,28 @@
is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z),
zz)"
- satisfies_b :: "[i,i,i]=>i"
+definition
+ satisfies_b :: "[i,i,i]=>i" where
"satisfies_b(A) ==
\<lambda>x y. \<lambda>env \<in> list(A). bool_of_o (nth(x,env) = nth(y,env))"
- satisfies_is_b :: "[i=>o,i,i,i,i]=>o"
+definition
+ satisfies_is_b :: "[i=>o,i,i,i,i]=>o" where
--{*We simplify the formula to have just @{term nx} rather than
introducing @{term ny} with @{term "nx=ny"} *}
- "satisfies_is_b(M,A) ==
+ "satisfies_is_b(M,A) ==
\<lambda>x y zz. \<forall>lA[M]. is_list(M,A,lA) -->
is_lambda(M, lA,
\<lambda>env z. is_bool_of_o(M,
\<exists>nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z),
zz)"
-
- satisfies_c :: "[i,i,i,i,i]=>i"
+
+definition
+ satisfies_c :: "[i,i,i,i,i]=>i" where
"satisfies_c(A) == \<lambda>p q rp rq. \<lambda>env \<in> list(A). not(rp ` env and rq ` env)"
- satisfies_is_c :: "[i=>o,i,i,i,i,i]=>o"
+definition
+ satisfies_is_c :: "[i=>o,i,i,i,i,i]=>o" where
"satisfies_is_c(M,A,h) ==
\<lambda>p q zz. \<forall>lA[M]. is_list(M,A,lA) -->
is_lambda(M, lA, \<lambda>env z. \<exists>hp[M]. \<exists>hq[M].
@@ -233,11 +240,13 @@
(\<exists>pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)),
zz)"
- satisfies_d :: "[i,i,i]=>i"
+definition
+ satisfies_d :: "[i,i,i]=>i" where
"satisfies_d(A)
== \<lambda>p rp. \<lambda>env \<in> list(A). bool_of_o (\<forall>x\<in>A. rp ` (Cons(x,env)) = 1)"
- satisfies_is_d :: "[i=>o,i,i,i,i]=>o"
+definition
+ satisfies_is_d :: "[i=>o,i,i,i,i]=>o" where
"satisfies_is_d(M,A,h) ==
\<lambda>p zz. \<forall>lA[M]. is_list(M,A,lA) -->
is_lambda(M, lA,
@@ -249,10 +258,11 @@
z),
zz)"
- satisfies_MH :: "[i=>o,i,i,i,i]=>o"
+definition
+ satisfies_MH :: "[i=>o,i,i,i,i]=>o" where
--{*The variable @{term u} is unused, but gives @{term satisfies_MH}
the correct arity.*}
- "satisfies_MH ==
+ "satisfies_MH ==
\<lambda>M A u f z.
\<forall>fml[M]. is_formula(M,fml) -->
is_lambda (M, fml,
@@ -261,8 +271,9 @@
satisfies_is_c(M,A,f), satisfies_is_d(M,A,f)),
z)"
- is_satisfies :: "[i=>o,i,i,i]=>o"
- "is_satisfies(M,A) == is_formula_rec (M, satisfies_MH(M,A))"
+definition
+ is_satisfies :: "[i=>o,i,i,i]=>o" where
+ "is_satisfies(M,A) == is_formula_rec (M, satisfies_MH(M,A))"
text{*This lemma relates the fragments defined above to the original primitive
@@ -504,7 +515,8 @@
2 1 0
finite_ordinal(M,dp) & is_depth(M,p,dp) & successor(M,dp,sdp) &
fun_apply(M,h,sdp,hsdp) & fun_apply(M,hsdp,p,z) *)
-definition depth_apply_fm :: "[i,i,i]=>i"
+definition
+ depth_apply_fm :: "[i,i,i]=>i" where
"depth_apply_fm(h,p,z) ==
Exists(Exists(Exists(
And(finite_ordinal_fm(2),
@@ -547,8 +559,9 @@
is_nth(M,x,env,nx) & is_nth(M,y,env,ny) & nx \<in> ny, z),
zz) *)
-definition satisfies_is_a_fm :: "[i,i,i,i]=>i"
- "satisfies_is_a_fm(A,x,y,z) ==
+definition
+ satisfies_is_a_fm :: "[i,i,i,i]=>i" where
+ "satisfies_is_a_fm(A,x,y,z) ==
Forall(
Implies(is_list_fm(succ(A),0),
lambda_fm(
@@ -598,7 +611,8 @@
\<exists>nx[M]. is_nth(M,x,env,nx) & is_nth(M,y,env,nx), z),
zz) *)
-definition satisfies_is_b_fm :: "[i,i,i,i]=>i"
+definition
+ satisfies_is_b_fm :: "[i,i,i,i]=>i" where
"satisfies_is_b_fm(A,x,y,z) ==
Forall(
Implies(is_list_fm(succ(A),0),
@@ -647,7 +661,8 @@
(\<exists>pq[M]. is_and(M,hp,hq,pq) & is_not(M,pq,z)),
zz) *)
-definition satisfies_is_c_fm :: "[i,i,i,i,i]=>i"
+definition
+ satisfies_is_c_fm :: "[i,i,i,i,i]=>i" where
"satisfies_is_c_fm(A,h,p,q,zz) ==
Forall(
Implies(is_list_fm(succ(A),0),
@@ -700,7 +715,8 @@
z),
zz) *)
-definition satisfies_is_d_fm :: "[i,i,i,i]=>i"
+definition
+ satisfies_is_d_fm :: "[i,i,i,i]=>i" where
"satisfies_is_d_fm(A,h,p,zz) ==
Forall(
Implies(is_list_fm(succ(A),0),
@@ -754,7 +770,8 @@
satisfies_is_c(M,A,f), satisfies_is_d(M,A,f)),
zz) *)
-definition satisfies_MH_fm :: "[i,i,i,i]=>i"
+definition
+ satisfies_MH_fm :: "[i,i,i,i]=>i" where
"satisfies_MH_fm(A,u,f,zz) ==
Forall(
Implies(is_formula_fm(0),
--- a/src/ZF/Constructible/WF_absolute.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/WF_absolute.thy Fri Nov 17 02:20:03 2006 +0100
@@ -10,7 +10,7 @@
subsection{*Transitive closure without fixedpoints*}
definition
- rtrancl_alt :: "[i,i]=>i"
+ rtrancl_alt :: "[i,i]=>i" where
"rtrancl_alt(A,r) ==
{p \<in> A*A. \<exists>n\<in>nat. \<exists>f \<in> succ(n) -> A.
(\<exists>x y. p = <x,y> & f`0 = x & f`n = y) &
@@ -60,8 +60,7 @@
definition
-
- rtran_closure_mem :: "[i=>o,i,i,i] => o"
+ rtran_closure_mem :: "[i=>o,i,i,i] => o" where
--{*The property of belonging to @{text "rtran_closure(r)"}*}
"rtran_closure_mem(M,A,r,p) ==
\<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M].
@@ -74,12 +73,14 @@
fun_apply(M,f,j,fj) & successor(M,j,sj) &
fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"
- rtran_closure :: "[i=>o,i,i] => o"
+definition
+ rtran_closure :: "[i=>o,i,i] => o" where
"rtran_closure(M,r,s) ==
\<forall>A[M]. is_field(M,r,A) -->
(\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))"
- tran_closure :: "[i=>o,i,i] => o"
+definition
+ tran_closure :: "[i=>o,i,i] => o" where
"tran_closure(M,r,t) ==
\<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)"
--- a/src/ZF/Constructible/WFrec.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/WFrec.thy Fri Nov 17 02:20:03 2006 +0100
@@ -272,7 +272,7 @@
subsection{*Relativization of the ZF Predicate @{term is_recfun}*}
definition
- M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
+ M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where
"M_is_recfun(M,MH,r,a,f) ==
\<forall>z[M]. z \<in> f <->
(\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M].
@@ -280,11 +280,13 @@
pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
xa \<in> r & MH(x, f_r_sx, y))"
- is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
+definition
+ is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o" where
"is_wfrec(M,MH,r,a,z) ==
\<exists>f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)"
- wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o"
+definition
+ wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where
"wfrec_replacement(M,MH,r) ==
strong_replacement(M,
\<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))"
--- a/src/ZF/Constructible/Wellorderings.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/Constructible/Wellorderings.thy Fri Nov 17 02:20:03 2006 +0100
@@ -17,28 +17,33 @@
subsection{*Wellorderings*}
definition
- irreflexive :: "[i=>o,i,i]=>o"
+ irreflexive :: "[i=>o,i,i]=>o" where
"irreflexive(M,A,r) == \<forall>x[M]. x\<in>A --> <x,x> \<notin> r"
- transitive_rel :: "[i=>o,i,i]=>o"
+definition
+ transitive_rel :: "[i=>o,i,i]=>o" where
"transitive_rel(M,A,r) ==
\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> (\<forall>z[M]. z\<in>A -->
<x,y>\<in>r --> <y,z>\<in>r --> <x,z>\<in>r))"
- linear_rel :: "[i=>o,i,i]=>o"
+definition
+ linear_rel :: "[i=>o,i,i]=>o" where
"linear_rel(M,A,r) ==
\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> <x,y>\<in>r | x=y | <y,x>\<in>r)"
- wellfounded :: "[i=>o,i]=>o"
+definition
+ wellfounded :: "[i=>o,i]=>o" where
--{*EVERY non-empty set has an @{text r}-minimal element*}
"wellfounded(M,r) ==
\<forall>x[M]. x\<noteq>0 --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
- wellfounded_on :: "[i=>o,i,i]=>o"
+definition
+ wellfounded_on :: "[i=>o,i,i]=>o" where
--{*every non-empty SUBSET OF @{text A} has an @{text r}-minimal element*}
"wellfounded_on(M,A,r) ==
\<forall>x[M]. x\<noteq>0 --> x\<subseteq>A --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
- wellordered :: "[i=>o,i,i]=>o"
+definition
+ wellordered :: "[i=>o,i,i]=>o" where
--{*linear and wellfounded on @{text A}*}
"wellordered(M,A,r) ==
transitive_rel(M,A,r) & linear_rel(M,A,r) & wellfounded_on(M,A,r)"
--- a/src/ZF/IMP/Denotation.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/IMP/Denotation.thy Fri Nov 17 02:20:03 2006 +0100
@@ -15,7 +15,7 @@
C :: "i => i"
definition
- Gamma :: "[i,i,i] => i" ("\<Gamma>")
+ Gamma :: "[i,i,i] => i" ("\<Gamma>") where
"\<Gamma>(b,cden) ==
(\<lambda>phi. {io \<in> (phi O cden). B(b,fst(io))=1} \<union>
{io \<in> id(loc->nat). B(b,fst(io))=0})"
--- a/src/ZF/ex/Commutation.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/ex/Commutation.thy Fri Nov 17 02:20:03 2006 +0100
@@ -10,23 +10,29 @@
theory Commutation imports Main begin
definition
- square :: "[i, i, i, i] => o"
+ square :: "[i, i, i, i] => o" where
"square(r,s,t,u) ==
(\<forall>a b. <a,b> \<in> r --> (\<forall>c. <a, c> \<in> s --> (\<exists>x. <b,x> \<in> t & <c,x> \<in> u)))"
- commute :: "[i, i] => o"
+definition
+ commute :: "[i, i] => o" where
"commute(r,s) == square(r,s,s,r)"
- diamond :: "i=>o"
+definition
+ diamond :: "i=>o" where
"diamond(r) == commute(r, r)"
- strip :: "i=>o"
+definition
+ strip :: "i=>o" where
"strip(r) == commute(r^*, r)"
- Church_Rosser :: "i => o"
+definition
+ Church_Rosser :: "i => o" where
"Church_Rosser(r) == (\<forall>x y. <x,y> \<in> (r Un converse(r))^* -->
(\<exists>z. <x,z> \<in> r^* & <y,z> \<in> r^*))"
- confluent :: "i=>o"
+
+definition
+ confluent :: "i=>o" where
"confluent(r) == diamond(r^*)"
--- a/src/ZF/ex/Group.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/ex/Group.thy Fri Nov 17 02:20:03 2006 +0100
@@ -21,20 +21,23 @@
*)
definition
- carrier :: "i => i"
+ carrier :: "i => i" where
"carrier(M) == fst(M)"
- mmult :: "[i, i, i] => i" (infixl "\<cdot>\<index>" 70)
+definition
+ mmult :: "[i, i, i] => i" (infixl "\<cdot>\<index>" 70) where
"mmult(M,x,y) == fst(snd(M)) ` <x,y>"
- one :: "i => i" ("\<one>\<index>")
+definition
+ one :: "i => i" ("\<one>\<index>") where
"one(M) == fst(snd(snd(M)))"
- update_carrier :: "[i,i] => i"
+definition
+ update_carrier :: "[i,i] => i" where
"update_carrier(M,A) == <A,snd(M)>"
definition
- m_inv :: "i => i => i" ("inv\<index> _" [81] 80)
+ m_inv :: "i => i => i" ("inv\<index> _" [81] 80) where
"inv\<^bsub>G\<^esub> x == (THE y. y \<in> carrier(G) & y \<cdot>\<^bsub>G\<^esub> x = \<one>\<^bsub>G\<^esub> & x \<cdot>\<^bsub>G\<^esub> y = \<one>\<^bsub>G\<^esub>)"
locale monoid = struct G +
@@ -295,7 +298,7 @@
subsection {* Direct Products *}
definition
- DirProdGroup :: "[i,i] => i" (infixr "\<Otimes>" 80)
+ DirProdGroup :: "[i,i] => i" (infixr "\<Otimes>" 80) where
"G \<Otimes> H == <carrier(G) \<times> carrier(H),
(\<lambda><<g,h>, <g', h'>>
\<in> (carrier(G) \<times> carrier(H)) \<times> (carrier(G) \<times> carrier(H)).
@@ -333,7 +336,7 @@
subsection {* Isomorphisms *}
definition
- hom :: "[i,i] => i"
+ hom :: "[i,i] => i" where
"hom(G,H) ==
{h \<in> carrier(G) -> carrier(H).
(\<forall>x \<in> carrier(G). \<forall>y \<in> carrier(G). h ` (x \<cdot>\<^bsub>G\<^esub> y) = (h ` x) \<cdot>\<^bsub>H\<^esub> (h ` y))}"
@@ -359,7 +362,7 @@
subsection {* Isomorphisms *}
definition
- iso :: "[i,i] => i" (infixr "\<cong>" 60)
+ iso :: "[i,i] => i" (infixr "\<cong>" 60) where
"G \<cong> H == hom(G,H) \<inter> bij(carrier(G), carrier(H))"
lemma (in group) iso_refl: "id(carrier(G)) \<in> G \<cong> G"
@@ -479,7 +482,7 @@
subsection {* Bijections of a Set, Permutation Groups, Automorphism Groups *}
definition
- BijGroup :: "i=>i"
+ BijGroup :: "i=>i" where
"BijGroup(S) ==
<bij(S,S),
\<lambda><g,f> \<in> bij(S,S) \<times> bij(S,S). g O f,
@@ -514,10 +517,11 @@
definition
- auto :: "i=>i"
+ auto :: "i=>i" where
"auto(G) == iso(G,G)"
- AutoGroup :: "i=>i"
+definition
+ AutoGroup :: "i=>i" where
"AutoGroup(G) == update_carrier(BijGroup(carrier(G)), auto(G))"
@@ -552,19 +556,23 @@
subsection{*Cosets and Quotient Groups*}
definition
- r_coset :: "[i,i,i] => i" (infixl "#>\<index>" 60)
+ r_coset :: "[i,i,i] => i" (infixl "#>\<index>" 60) where
"H #>\<^bsub>G\<^esub> a == \<Union>h\<in>H. {h \<cdot>\<^bsub>G\<^esub> a}"
- l_coset :: "[i,i,i] => i" (infixl "<#\<index>" 60)
+definition
+ l_coset :: "[i,i,i] => i" (infixl "<#\<index>" 60) where
"a <#\<^bsub>G\<^esub> H == \<Union>h\<in>H. {a \<cdot>\<^bsub>G\<^esub> h}"
- RCOSETS :: "[i,i] => i" ("rcosets\<index> _" [81] 80)
+definition
+ RCOSETS :: "[i,i] => i" ("rcosets\<index> _" [81] 80) where
"rcosets\<^bsub>G\<^esub> H == \<Union>a\<in>carrier(G). {H #>\<^bsub>G\<^esub> a}"
- set_mult :: "[i,i,i] => i" (infixl "<#>\<index>" 60)
+definition
+ set_mult :: "[i,i,i] => i" (infixl "<#>\<index>" 60) where
"H <#>\<^bsub>G\<^esub> K == \<Union>h\<in>H. \<Union>k\<in>K. {h \<cdot>\<^bsub>G\<^esub> k}"
- SET_INV :: "[i,i] => i" ("set'_inv\<index> _" [81] 80)
+definition
+ SET_INV :: "[i,i] => i" ("set'_inv\<index> _" [81] 80) where
"set_inv\<^bsub>G\<^esub> H == \<Union>h\<in>H. {inv\<^bsub>G\<^esub> h}"
@@ -833,7 +841,7 @@
subsubsection{*Two distinct right cosets are disjoint*}
definition
- r_congruent :: "[i,i] => i" ("rcong\<index> _" [60] 60)
+ r_congruent :: "[i,i] => i" ("rcong\<index> _" [60] 60) where
"rcong\<^bsub>G\<^esub> H == {<x,y> \<in> carrier(G) * carrier(G). inv\<^bsub>G\<^esub> x \<cdot>\<^bsub>G\<^esub> y \<in> H}"
@@ -900,7 +908,7 @@
subsection {*Order of a Group and Lagrange's Theorem*}
definition
- order :: "i => i"
+ order :: "i => i" where
"order(S) == |carrier(S)|"
lemma (in group) rcos_self:
@@ -972,7 +980,7 @@
subsection {*Quotient Groups: Factorization of a Group*}
definition
- FactGroup :: "[i,i] => i" (infixl "Mod" 65)
+ FactGroup :: "[i,i] => i" (infixl "Mod" 65) where
--{*Actually defined for groups rather than monoids*}
"G Mod H ==
<rcosets\<^bsub>G\<^esub> H, \<lambda><K1,K2> \<in> (rcosets\<^bsub>G\<^esub> H) \<times> (rcosets\<^bsub>G\<^esub> H). K1 <#>\<^bsub>G\<^esub> K2, H, 0>"
@@ -1035,7 +1043,7 @@
range of that homomorphism.*}
definition
- kernel :: "[i,i,i] => i"
+ kernel :: "[i,i,i] => i" where
--{*the kernel of a homomorphism*}
"kernel(G,H,h) == {x \<in> carrier(G). h ` x = \<one>\<^bsub>H\<^esub>}";
--- a/src/ZF/ex/Ramsey.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/ex/Ramsey.thy Fri Nov 17 02:20:03 2006 +0100
@@ -29,19 +29,23 @@
theory Ramsey imports Main begin
definition
- Symmetric :: "i=>o"
+ Symmetric :: "i=>o" where
"Symmetric(E) == (\<forall>x y. <x,y>:E --> <y,x>:E)"
- Atleast :: "[i,i]=>o" (*not really necessary: ZF defines cardinality*)
+definition
+ Atleast :: "[i,i]=>o" where -- "not really necessary: ZF defines cardinality"
"Atleast(n,S) == (\<exists>f. f \<in> inj(n,S))"
- Clique :: "[i,i,i]=>o"
+definition
+ Clique :: "[i,i,i]=>o" where
"Clique(C,V,E) == (C \<subseteq> V) & (\<forall>x \<in> C. \<forall>y \<in> C. x\<noteq>y --> <x,y> \<in> E)"
- Indept :: "[i,i,i]=>o"
+definition
+ Indept :: "[i,i,i]=>o" where
"Indept(I,V,E) == (I \<subseteq> V) & (\<forall>x \<in> I. \<forall>y \<in> I. x\<noteq>y --> <x,y> \<notin> E)"
- Ramsey :: "[i,i,i]=>o"
+definition
+ Ramsey :: "[i,i,i]=>o" where
"Ramsey(n,i,j) == \<forall>V E. Symmetric(E) & Atleast(n,V) -->
(\<exists>C. Clique(C,V,E) & Atleast(i,C)) |
(\<exists>I. Indept(I,V,E) & Atleast(j,I))"
--- a/src/ZF/ex/Ring.thy Fri Nov 17 02:19:55 2006 +0100
+++ b/src/ZF/ex/Ring.thy Fri Nov 17 02:20:03 2006 +0100
@@ -14,13 +14,15 @@
*)
definition
- add_field :: "i => i"
+ add_field :: "i => i" where
"add_field(M) = fst(snd(snd(snd(M))))"
- ring_add :: "[i, i, i] => i" (infixl "\<oplus>\<index>" 65)
+definition
+ ring_add :: "[i, i, i] => i" (infixl "\<oplus>\<index>" 65) where
"ring_add(M,x,y) = add_field(M) ` <x,y>"
- zero :: "i => i" ("\<zero>\<index>")
+definition
+ zero :: "i => i" ("\<zero>\<index>") where
"zero(M) = fst(snd(snd(snd(snd(M)))))"
@@ -37,10 +39,11 @@
text {* Derived operations. *}
definition
- a_inv :: "[i,i] => i" ("\<ominus>\<index> _" [81] 80)
+ a_inv :: "[i,i] => i" ("\<ominus>\<index> _" [81] 80) where
"a_inv(R) == m_inv (<carrier(R), add_field(R), zero(R), 0>)"
- minus :: "[i,i,i] => i" (infixl "\<ominus>\<index>" 65)
+definition
+ minus :: "[i,i,i] => i" (infixl "\<ominus>\<index>" 65) where
"\<lbrakk>x \<in> carrier(R); y \<in> carrier(R)\<rbrakk> \<Longrightarrow> x \<ominus>\<^bsub>R\<^esub> y = x \<oplus>\<^bsub>R\<^esub> (\<ominus>\<^bsub>R\<^esub> y)"
locale abelian_monoid = struct G +
@@ -225,7 +228,7 @@
subsection {* Morphisms *}
definition
- ring_hom :: "[i,i] => i"
+ ring_hom :: "[i,i] => i" where
"ring_hom(R,S) ==
{h \<in> carrier(R) -> carrier(S).
(\<forall>x y. x \<in> carrier(R) & y \<in> carrier(R) -->