# HG changeset patch # User haftmann # Date 1236334257 -3600 # Node ID 6c74ef5a349f51eb23311dce44a078250d86b4da # Parent d6bffd97d8d56e1f3c3b25530f31dccc3721444a# Parent 8f4d5eaa98789e18e746fbe27659f247d3a7c9ff merged diff -r d6bffd97d8d5 -r 6c74ef5a349f NEWS --- a/NEWS Fri Mar 06 09:35:43 2009 +0100 +++ b/NEWS Fri Mar 06 11:10:57 2009 +0100 @@ -220,6 +220,21 @@ *** HOL *** +* Some set operations are now proper qualified constants with authentic syntax. +INCOMPATIBILITY: + + op Int ~> Set.Int + op Un ~> Set.Un + INTER ~> Set.INTER + UNION ~> Set.UNION + Inter ~> Set.Inter + Union ~> Set.Union + {} ~> Set.empty + UNIV ~> Set.UNIV + +* Class complete_lattice with operations Inf, Sup, INFI, SUPR now in theory +Set. + * Auxiliary class "itself" has disappeared -- classes without any parameter are treated as expected by the 'class' command. @@ -297,7 +312,7 @@ avoiding strange error messages. * Simplifier: simproc for let expressions now unfolds if bound variable -occurs at most one time in let expression body. INCOMPATIBILITY. +occurs at most once in let expression body. INCOMPATIBILITY. * New classes "top" and "bot" with corresponding operations "top" and "bot" in theory Orderings; instantiation of class "complete_lattice" requires diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Fun.thy --- a/src/HOL/Fun.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Fun.thy Fri Mar 06 11:10:57 2009 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Fun.thy - ID: $Id$ Author: Tobias Nipkow, Cambridge University Computer Laboratory Copyright 1994 University of Cambridge *) diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Hoare/Hoare.thy --- a/src/HOL/Hoare/Hoare.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Hoare/Hoare.thy Fri Mar 06 11:10:57 2009 +0100 @@ -161,9 +161,9 @@ (* assn_tr' & bexp_tr'*) ML{* fun assn_tr' (Const ("Collect",_) $ T) = dest_abstuple T - | assn_tr' (Const ("op Int",_) $ (Const ("Collect",_) $ T1) $ + | assn_tr' (Const (@{const_name Set.Int}, _) $ (Const ("Collect",_) $ T1) $ (Const ("Collect",_) $ T2)) = - Syntax.const "op Int" $ dest_abstuple T1 $ dest_abstuple T2 + Syntax.const "Set.Int" $ dest_abstuple T1 $ dest_abstuple T2 | assn_tr' t = t; fun bexp_tr' (Const ("Collect",_) $ T) = dest_abstuple T diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Hoare/HoareAbort.thy --- a/src/HOL/Hoare/HoareAbort.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Hoare/HoareAbort.thy Fri Mar 06 11:10:57 2009 +0100 @@ -163,9 +163,9 @@ (* assn_tr' & bexp_tr'*) ML{* fun assn_tr' (Const ("Collect",_) $ T) = dest_abstuple T - | assn_tr' (Const ("op Int",_) $ (Const ("Collect",_) $ T1) $ + | assn_tr' (Const (@{const_name Int},_) $ (Const ("Collect",_) $ T1) $ (Const ("Collect",_) $ T2)) = - Syntax.const "op Int" $ dest_abstuple T1 $ dest_abstuple T2 + Syntax.const "Set.Int" $ dest_abstuple T1 $ dest_abstuple T2 | assn_tr' t = t; fun bexp_tr' (Const ("Collect",_) $ T) = dest_abstuple T diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/HoareParallel/OG_Tran.thy --- a/src/HOL/HoareParallel/OG_Tran.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/HoareParallel/OG_Tran.thy Fri Mar 06 11:10:57 2009 +0100 @@ -102,7 +102,7 @@ "SEM c S \ \sem c ` S " syntax "_Omega" :: "'a com" ("\" 63) -translations "\" \ "While UNIV UNIV (Basic id)" +translations "\" \ "While CONST UNIV CONST UNIV (Basic id)" consts fwhile :: "'a bexp \ 'a com \ nat \ 'a com" primrec diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Lattices.thy --- a/src/HOL/Lattices.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Lattices.thy Fri Mar 06 11:10:57 2009 +0100 @@ -5,7 +5,7 @@ header {* Abstract lattices *} theory Lattices -imports Fun +imports Orderings begin subsection {* Lattices *} @@ -328,135 +328,6 @@ min_max.le_infI1 min_max.le_infI2 -subsection {* Complete lattices *} - -class complete_lattice = lattice + bot + top + - fixes Inf :: "'a set \ 'a" ("\_" [900] 900) - and Sup :: "'a set \ 'a" ("\_" [900] 900) - assumes Inf_lower: "x \ A \ \A \ x" - and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" - assumes Sup_upper: "x \ A \ x \ \A" - and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" -begin - -lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" - by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) - -lemma Inf_Univ: "\UNIV = \{}" - unfolding Sup_Inf by auto - -lemma Sup_Univ: "\UNIV = \{}" - unfolding Inf_Sup by auto - -lemma Inf_insert: "\insert a A = a \ \A" - by (auto intro: antisym Inf_greatest Inf_lower) - -lemma Sup_insert: "\insert a A = a \ \A" - by (auto intro: antisym Sup_least Sup_upper) - -lemma Inf_singleton [simp]: - "\{a} = a" - by (auto intro: antisym Inf_lower Inf_greatest) - -lemma Sup_singleton [simp]: - "\{a} = a" - by (auto intro: antisym Sup_upper Sup_least) - -lemma Inf_insert_simp: - "\insert a A = (if A = {} then a else a \ \A)" - by (cases "A = {}") (simp_all, simp add: Inf_insert) - -lemma Sup_insert_simp: - "\insert a A = (if A = {} then a else a \ \A)" - by (cases "A = {}") (simp_all, simp add: Sup_insert) - -lemma Inf_binary: - "\{a, b} = a \ b" - by (simp add: Inf_insert_simp) - -lemma Sup_binary: - "\{a, b} = a \ b" - by (simp add: Sup_insert_simp) - -lemma bot_def: - "bot = \{}" - by (auto intro: antisym Sup_least) - -lemma top_def: - "top = \{}" - by (auto intro: antisym Inf_greatest) - -lemma sup_bot [simp]: - "x \ bot = x" - using bot_least [of x] by (simp add: le_iff_sup sup_commute) - -lemma inf_top [simp]: - "x \ top = x" - using top_greatest [of x] by (simp add: le_iff_inf inf_commute) - -definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where - "SUPR A f == \ (f ` A)" - -definition INFI :: "'b set \ ('b \ 'a) \ 'a" where - "INFI A f == \ (f ` A)" - -end - -syntax - "_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10) - "_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 10] 10) - "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) - "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10) - -translations - "SUP x y. B" == "SUP x. SUP y. B" - "SUP x. B" == "CONST SUPR UNIV (%x. B)" - "SUP x. B" == "SUP x:UNIV. B" - "SUP x:A. B" == "CONST SUPR A (%x. B)" - "INF x y. B" == "INF x. INF y. B" - "INF x. B" == "CONST INFI UNIV (%x. B)" - "INF x. B" == "INF x:UNIV. B" - "INF x:A. B" == "CONST INFI A (%x. B)" - -(* To avoid eta-contraction of body: *) -print_translation {* -let - fun btr' syn (A :: Abs abs :: ts) = - let val (x,t) = atomic_abs_tr' abs - in list_comb (Syntax.const syn $ x $ A $ t, ts) end - val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const -in -[(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")] -end -*} - -context complete_lattice -begin - -lemma le_SUPI: "i : A \ M i \ (SUP i:A. M i)" - by (auto simp add: SUPR_def intro: Sup_upper) - -lemma SUP_leI: "(\i. i : A \ M i \ u) \ (SUP i:A. M i) \ u" - by (auto simp add: SUPR_def intro: Sup_least) - -lemma INF_leI: "i : A \ (INF i:A. M i) \ M i" - by (auto simp add: INFI_def intro: Inf_lower) - -lemma le_INFI: "(\i. i : A \ u \ M i) \ u \ (INF i:A. M i)" - by (auto simp add: INFI_def intro: Inf_greatest) - -lemma SUP_const[simp]: "A \ {} \ (SUP i:A. M) = M" - by (auto intro: antisym SUP_leI le_SUPI) - -lemma INF_const[simp]: "A \ {} \ (INF i:A. M) = M" - by (auto intro: antisym INF_leI le_INFI) - -end - - subsection {* Bool as lattice *} instantiation bool :: distrib_lattice @@ -473,28 +344,6 @@ end -instantiation bool :: complete_lattice -begin - -definition - Inf_bool_def: "\A \ (\x\A. x)" - -definition - Sup_bool_def: "\A \ (\x\A. x)" - -instance - by intro_classes (auto simp add: Inf_bool_def Sup_bool_def le_bool_def) - -end - -lemma Inf_empty_bool [simp]: - "\{}" - unfolding Inf_bool_def by auto - -lemma not_Sup_empty_bool [simp]: - "\ Sup {}" - unfolding Sup_bool_def by auto - subsection {* Fun as lattice *} @@ -522,85 +371,6 @@ instance "fun" :: (type, distrib_lattice) distrib_lattice by default (auto simp add: inf_fun_eq sup_fun_eq sup_inf_distrib1) -instantiation "fun" :: (type, complete_lattice) complete_lattice -begin - -definition - Inf_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" - -definition - Sup_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" - -instance - by intro_classes - (auto simp add: Inf_fun_def Sup_fun_def le_fun_def - intro: Inf_lower Sup_upper Inf_greatest Sup_least) - -end - -lemma Inf_empty_fun: - "\{} = (\_. \{})" - by rule (auto simp add: Inf_fun_def) - -lemma Sup_empty_fun: - "\{} = (\_. \{})" - by rule (auto simp add: Sup_fun_def) - - -subsection {* Set as lattice *} - -lemma inf_set_eq: "A \ B = A \ B" - apply (rule subset_antisym) - apply (rule Int_greatest) - apply (rule inf_le1) - apply (rule inf_le2) - apply (rule inf_greatest) - apply (rule Int_lower1) - apply (rule Int_lower2) - done - -lemma sup_set_eq: "A \ B = A \ B" - apply (rule subset_antisym) - apply (rule sup_least) - apply (rule Un_upper1) - apply (rule Un_upper2) - apply (rule Un_least) - apply (rule sup_ge1) - apply (rule sup_ge2) - done - -lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" - apply (fold inf_set_eq sup_set_eq) - apply (erule mono_inf) - done - -lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" - apply (fold inf_set_eq sup_set_eq) - apply (erule mono_sup) - done - -lemma Inf_set_eq: "\S = \S" - apply (rule subset_antisym) - apply (rule Inter_greatest) - apply (erule Inf_lower) - apply (rule Inf_greatest) - apply (erule Inter_lower) - done - -lemma Sup_set_eq: "\S = \S" - apply (rule subset_antisym) - apply (rule Sup_least) - apply (erule Union_upper) - apply (rule Union_least) - apply (erule Sup_upper) - done - -lemma top_set_eq: "top = UNIV" - by (iprover intro!: subset_antisym subset_UNIV top_greatest) - -lemma bot_set_eq: "bot = {}" - by (iprover intro!: subset_antisym empty_subsetI bot_least) - text {* redundant bindings *} @@ -611,8 +381,6 @@ less_eq (infix "\" 50) and less (infix "\" 50) and inf (infixl "\" 70) and - sup (infixl "\" 65) and - Inf ("\_" [900] 900) and - Sup ("\_" [900] 900) + sup (infixl "\" 65) end diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Library/Euclidean_Space.thy --- a/src/HOL/Library/Euclidean_Space.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Library/Euclidean_Space.thy Fri Mar 06 11:10:57 2009 +0100 @@ -4199,7 +4199,7 @@ assumes iB: "independent (B:: (real ^'n) set)" shows "\g. linear g \ (\x\B. g x = f x)" proof- - from maximal_independent_subset_extend[of B "UNIV"] iB + from maximal_independent_subset_extend[of B UNIV] iB obtain C where C: "B \ C" "independent C" "\x. x \ span C" by auto from C(2) independent_bound[of C] linear_independent_extend_lemma[of C f] diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Library/Executable_Set.thy --- a/src/HOL/Library/Executable_Set.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Library/Executable_Set.thy Fri Mar 06 11:10:57 2009 +0100 @@ -178,7 +178,7 @@ lemma iso_insert: "set (insertl x xs) = insert x (set xs)" - unfolding insertl_def iso_member by (simp add: Set.insert_absorb) + unfolding insertl_def iso_member by (simp add: insert_absorb) lemma iso_remove1: assumes distnct: "distinct xs" @@ -261,7 +261,7 @@ subsubsection {* const serializations *} consts_code - "{}" ("{*[]*}") + "Set.empty" ("{*[]*}") insert ("{*insertl*}") is_empty ("{*null*}") "op \" ("{*unionl*}") diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Library/Finite_Cartesian_Product.thy --- a/src/HOL/Library/Finite_Cartesian_Product.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Library/Finite_Cartesian_Product.thy Fri Mar 06 11:10:57 2009 +0100 @@ -15,7 +15,7 @@ definition "dimindex (S:: 'a set) = (if finite (UNIV::'a set) then card (UNIV:: 'a set) else 1)" syntax "_type_dimindex" :: "type => nat" ("(1DIM/(1'(_')))") -translations "DIM(t)" => "CONST dimindex (UNIV :: t set)" +translations "DIM(t)" => "CONST dimindex (CONST UNIV :: t set)" lemma dimindex_nonzero: "dimindex S \ 0" unfolding dimindex_def diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Library/LaTeXsugar.thy --- a/src/HOL/Library/LaTeXsugar.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Library/LaTeXsugar.thy Fri Mar 06 11:10:57 2009 +0100 @@ -30,7 +30,7 @@ (* empty set *) notation (latex) - "{}" ("\") + "Set.empty" ("\") (* insert *) translations diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Nominal/nominal_inductive2.ML Fri Mar 06 11:10:57 2009 +0100 @@ -75,7 +75,7 @@ | add_binders thy i (Abs (_, _, t)) bs = add_binders thy (i + 1) t bs | add_binders thy i _ bs = bs; -fun mk_set T [] = Const ("{}", HOLogic.mk_setT T) +fun mk_set T [] = Const (@{const_name Set.empty}, HOLogic.mk_setT T) | mk_set T (x :: xs) = Const ("insert", T --> HOLogic.mk_setT T --> HOLogic.mk_setT T) $ x $ mk_set T xs; @@ -201,7 +201,7 @@ | add_set (t, T) ((u, U) :: ps) = if T = U then let val S = HOLogic.mk_setT T - in (Const (@{const_name "op Un"}, S --> S --> S) $ u $ t, T) :: ps + in (Const (@{const_name "Un"}, S --> S --> S) $ u $ t, T) :: ps end else (u, U) :: add_set (t, T) ps in diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Nominal/nominal_package.ML --- a/src/HOL/Nominal/nominal_package.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Nominal/nominal_package.ML Fri Mar 06 11:10:57 2009 +0100 @@ -1010,8 +1010,8 @@ (augment_sort thy8 pt_cp_sort (HOLogic.mk_Trueprop (HOLogic.mk_eq (supp c, - if null dts then Const ("{}", HOLogic.mk_setT atomT) - else foldr1 (HOLogic.mk_binop "op Un") (map supp args2))))) + if null dts then Const (@{const_name Set.empty}, HOLogic.mk_setT atomT) + else foldr1 (HOLogic.mk_binop @{const_name Un}) (map supp args2))))) (fn _ => simp_tac (HOL_basic_ss addsimps (supp_def :: Un_assoc :: de_Morgan_conj :: Collect_disj_eq :: finite_Un :: diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Set.thy --- a/src/HOL/Set.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Set.thy Fri Mar 06 11:10:57 2009 +0100 @@ -1,12 +1,11 @@ (* Title: HOL/Set.thy - ID: $Id$ Author: Tobias Nipkow, Lawrence C Paulson and Markus Wenzel *) header {* Set theory for higher-order logic *} theory Set -imports Orderings +imports Lattices begin text {* A set in HOL is simply a predicate. *} @@ -19,36 +18,21 @@ types 'a set = "'a => bool" consts - "{}" :: "'a set" ("{}") - UNIV :: "'a set" + Collect :: "('a => bool) => 'a set" -- "comprehension" + "op :" :: "'a => 'a set => bool" -- "membership" insert :: "'a => 'a set => 'a set" - Collect :: "('a => bool) => 'a set" -- "comprehension" - "op Int" :: "'a set => 'a set => 'a set" (infixl "Int" 70) - "op Un" :: "'a set => 'a set => 'a set" (infixl "Un" 65) - UNION :: "'a set => ('a => 'b set) => 'b set" -- "general union" - INTER :: "'a set => ('a => 'b set) => 'b set" -- "general intersection" - Union :: "'a set set => 'a set" -- "union of a set" - Inter :: "'a set set => 'a set" -- "intersection of a set" - Pow :: "'a set => 'a set set" -- "powerset" Ball :: "'a set => ('a => bool) => bool" -- "bounded universal quantifiers" Bex :: "'a set => ('a => bool) => bool" -- "bounded existential quantifiers" Bex1 :: "'a set => ('a => bool) => bool" -- "bounded unique existential quantifiers" + Pow :: "'a set => 'a set set" -- "powerset" image :: "('a => 'b) => 'a set => 'b set" (infixr "`" 90) - "op :" :: "'a => 'a set => bool" -- "membership" + +local notation "op :" ("op :") and "op :" ("(_/ : _)" [50, 51] 50) -local - - -subsection {* Additional concrete syntax *} - -abbreviation - range :: "('a => 'b) => 'b set" where -- "of function" - "range f == f ` UNIV" - abbreviation "not_mem x A == ~ (x : A)" -- "non-membership" @@ -57,32 +41,51 @@ not_mem ("(_/ ~: _)" [50, 51] 50) notation (xsymbols) - "op Int" (infixl "\" 70) and - "op Un" (infixl "\" 65) and "op :" ("op \") and "op :" ("(_/ \ _)" [50, 51] 50) and not_mem ("op \") and - not_mem ("(_/ \ _)" [50, 51] 50) and - Union ("\_" [90] 90) and - Inter ("\_" [90] 90) + not_mem ("(_/ \ _)" [50, 51] 50) notation (HTML output) - "op Int" (infixl "\" 70) and - "op Un" (infixl "\" 65) and "op :" ("op \") and "op :" ("(_/ \ _)" [50, 51] 50) and not_mem ("op \") and not_mem ("(_/ \ _)" [50, 51] 50) syntax + "@Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})") + +translations + "{x. P}" == "Collect (%x. P)" + +definition empty :: "'a set" ("{}") where + "empty \ {x. False}" + +definition UNIV :: "'a set" where + "UNIV \ {x. True}" + +syntax "@Finset" :: "args => 'a set" ("{(_)}") - "@Coll" :: "pttrn => bool => 'a set" ("(1{_./ _})") - "@SetCompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})") - "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ :/ _./ _})") - "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) - "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) - "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 10] 10) - "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 10] 10) + +translations + "{x, xs}" == "insert x {xs}" + "{x}" == "insert x {}" + +definition Int :: "'a set \ 'a set \ 'a set" (infixl "Int" 70) where + "A Int B \ {x. x \ A \ x \ B}" + +definition Un :: "'a set \ 'a set \ 'a set" (infixl "Un" 65) where + "A Un B \ {x. x \ A \ x \ B}" + +notation (xsymbols) + "Int" (infixl "\" 70) and + "Un" (infixl "\" 65) + +notation (HTML output) + "Int" (infixl "\" 70) and + "Un" (infixl "\" 65) + +syntax "_Ball" :: "pttrn => 'a set => bool => bool" ("(3ALL _:_./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn => 'a set => bool => bool" ("(3EX _:_./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3EX! _:_./ _)" [0, 0, 10] 10) @@ -93,24 +96,6 @@ "_Bex" :: "pttrn => 'a set => bool => bool" ("(3? _:_./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3?! _:_./ _)" [0, 0, 10] 10) -translations - "{x, xs}" == "insert x {xs}" - "{x}" == "insert x {}" - "{x. P}" == "Collect (%x. P)" - "{x:A. P}" => "{x. x:A & P}" - "UN x y. B" == "UN x. UN y. B" - "UN x. B" == "UNION UNIV (%x. B)" - "UN x. B" == "UN x:UNIV. B" - "INT x y. B" == "INT x. INT y. B" - "INT x. B" == "INTER UNIV (%x. B)" - "INT x. B" == "INT x:UNIV. B" - "UN x:A. B" == "UNION A (%x. B)" - "INT x:A. B" == "INTER A (%x. B)" - "ALL x:A. P" == "Ball A (%x. P)" - "EX x:A. P" == "Bex A (%x. P)" - "EX! x:A. P" == "Bex1 A (%x. P)" - "LEAST x:A. P" => "LEAST x. x:A & P" - syntax (xsymbols) "_Ball" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) @@ -122,26 +107,71 @@ "_Bex" :: "pttrn => 'a set => bool => bool" ("(3\_\_./ _)" [0, 0, 10] 10) "_Bex1" :: "pttrn => 'a set => bool => bool" ("(3\!_\_./ _)" [0, 0, 10] 10) +translations + "ALL x:A. P" == "Ball A (%x. P)" + "EX x:A. P" == "Bex A (%x. P)" + "EX! x:A. P" == "Bex1 A (%x. P)" + "LEAST x:A. P" => "LEAST x. x:A & P" + +definition INTER :: "'a set \ ('a \ 'b set) \ 'b set" where + "INTER A B \ {y. \x\A. y \ B x}" + +definition UNION :: "'a set \ ('a \ 'b set) \ 'b set" where + "UNION A B \ {y. \x\A. y \ B x}" + +definition Inter :: "'a set set \ 'a set" where + "Inter S \ INTER S (\x. x)" + +definition Union :: "'a set set \ 'a set" where + "Union S \ UNION S (\x. x)" + +notation (xsymbols) + Inter ("\_" [90] 90) and + Union ("\_" [90] 90) + + +subsection {* Additional concrete syntax *} + +syntax + "@SetCompr" :: "'a => idts => bool => 'a set" ("(1{_ |/_./ _})") + "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ :/ _./ _})") + "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) + "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) + "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3INT _:_./ _)" [0, 10] 10) + "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3UN _:_./ _)" [0, 10] 10) + syntax (xsymbols) "@Collect" :: "idt => 'a set => bool => 'a set" ("(1{_ \/ _./ _})") + "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) - "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\_./ _)" [0, 10] 10) + "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) - "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\_\_./ _)" [0, 10] 10) syntax (latex output) + "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) "@UNION1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) - "@INTER1" :: "pttrns => 'b set => 'b set" ("(3\(00\<^bsub>_\<^esub>)/ _)" [0, 10] 10) + "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) "@UNION" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) - "@INTER" :: "pttrn => 'a set => 'b set => 'b set" ("(3\(00\<^bsub>_\_\<^esub>)/ _)" [0, 10] 10) - -text{* + +translations + "{x:A. P}" => "{x. x:A & P}" + "INT x y. B" == "INT x. INT y. B" + "INT x. B" == "CONST INTER CONST UNIV (%x. B)" + "INT x. B" == "INT x:CONST UNIV. B" + "INT x:A. B" == "CONST INTER A (%x. B)" + "UN x y. B" == "UN x. UN y. B" + "UN x. B" == "CONST UNION CONST UNIV (%x. B)" + "UN x. B" == "UN x:CONST UNIV. B" + "UN x:A. B" == "CONST UNION A (%x. B)" + +text {* Note the difference between ordinary xsymbol syntax of indexed unions and intersections (e.g.\ @{text"\a\<^isub>1\A\<^isub>1. B"}) and their \LaTeX\ rendition: @{term"\a\<^isub>1\A\<^isub>1. B"}. The former does not make the index expression a subscript of the union/intersection symbol because this leads to problems with nested - subscripts in Proof General. *} + subscripts in Proof General. +*} abbreviation subset :: "'a set \ 'a set \ bool" where @@ -183,6 +213,10 @@ supset_eq ("op \") and supset_eq ("(_/ \ _)" [50, 51] 50) +abbreviation + range :: "('a => 'b) => 'b set" where -- "of function" + "range f == f ` UNIV" + subsubsection "Bounded quantifiers" @@ -280,12 +314,12 @@ (* To avoid eta-contraction of body: *) print_translation {* let - fun btr' syn [A,Abs abs] = - let val (x,t) = atomic_abs_tr' abs + fun btr' syn [A, Abs abs] = + let val (x, t) = atomic_abs_tr' abs in Syntax.const syn $ x $ A $ t end in -[("Ball", btr' "_Ball"),("Bex", btr' "_Bex"), - ("UNION", btr' "@UNION"),("INTER", btr' "@INTER")] +[(@{const_syntax Ball}, btr' "_Ball"), (@{const_syntax Bex}, btr' "_Bex"), + (@{const_syntax UNION}, btr' "@UNION"),(@{const_syntax INTER}, btr' "@INTER")] end *} @@ -373,15 +407,7 @@ end defs - Un_def: "A Un B == {x. x:A | x:B}" - Int_def: "A Int B == {x. x:A & x:B}" - INTER_def: "INTER A B == {y. ALL x:A. y: B(x)}" - UNION_def: "UNION A B == {y. EX x:A. y: B(x)}" - Inter_def: "Inter S == (INT x:S. x)" - Union_def: "Union S == (UN x:S. x)" Pow_def: "Pow A == {B. B <= A}" - empty_def: "{} == {x. False}" - UNIV_def: "UNIV == {x. True}" insert_def: "insert a B == {x. x=a} Un B" image_def: "f`A == {y. EX x:A. y = f(x)}" @@ -1048,12 +1074,12 @@ to use term-nets to associate patterns with rules. Also, if a rule fails to apply, then the formula should be kept. [("HOL.uminus", Compl_iff RS iffD1), ("HOL.minus", [Diff_iff RS iffD1]), - ("op Int", [IntD1,IntD2]), + ("Int", [IntD1,IntD2]), ("Collect", [CollectD]), ("Inter", [InterD]), ("INTER", [INT_D])] *) ML {* - val mksimps_pairs = [("Ball", @{thms bspec})] @ mksimps_pairs; + val mksimps_pairs = [(@{const_name Ball}, @{thms bspec})] @ mksimps_pairs; *} declaration {* fn _ => Simplifier.map_ss (fn ss => ss setmksimps (mksimps mksimps_pairs)) @@ -2160,9 +2186,7 @@ subsection {* Getting the Contents of a Singleton Set *} -definition - contents :: "'a set \ 'a" -where +definition contents :: "'a set \ 'a" where [code del]: "contents X = (THE x. X = {x})" lemma contents_eq [simp]: "contents {x} = x" @@ -2215,6 +2239,255 @@ unfolding vimage_def Collect_def mem_def .. +subsection {* Complete lattices *} + +notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + inf (infixl "\" 70) and + sup (infixl "\" 65) + +class complete_lattice = lattice + bot + top + + fixes Inf :: "'a set \ 'a" ("\_" [900] 900) + and Sup :: "'a set \ 'a" ("\_" [900] 900) + assumes Inf_lower: "x \ A \ \A \ x" + and Inf_greatest: "(\x. x \ A \ z \ x) \ z \ \A" + assumes Sup_upper: "x \ A \ x \ \A" + and Sup_least: "(\x. x \ A \ x \ z) \ \A \ z" +begin + +lemma Inf_Sup: "\A = \{b. \a \ A. b \ a}" + by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) + +lemma Sup_Inf: "\A = \{b. \a \ A. a \ b}" + by (auto intro: antisym Inf_lower Inf_greatest Sup_upper Sup_least) + +lemma Inf_Univ: "\UNIV = \{}" + unfolding Sup_Inf by auto + +lemma Sup_Univ: "\UNIV = \{}" + unfolding Inf_Sup by auto + +lemma Inf_insert: "\insert a A = a \ \A" + by (auto intro: antisym Inf_greatest Inf_lower) + +lemma Sup_insert: "\insert a A = a \ \A" + by (auto intro: antisym Sup_least Sup_upper) + +lemma Inf_singleton [simp]: + "\{a} = a" + by (auto intro: antisym Inf_lower Inf_greatest) + +lemma Sup_singleton [simp]: + "\{a} = a" + by (auto intro: antisym Sup_upper Sup_least) + +lemma Inf_insert_simp: + "\insert a A = (if A = {} then a else a \ \A)" + by (cases "A = {}") (simp_all, simp add: Inf_insert) + +lemma Sup_insert_simp: + "\insert a A = (if A = {} then a else a \ \A)" + by (cases "A = {}") (simp_all, simp add: Sup_insert) + +lemma Inf_binary: + "\{a, b} = a \ b" + by (simp add: Inf_insert_simp) + +lemma Sup_binary: + "\{a, b} = a \ b" + by (simp add: Sup_insert_simp) + +lemma bot_def: + "bot = \{}" + by (auto intro: antisym Sup_least) + +lemma top_def: + "top = \{}" + by (auto intro: antisym Inf_greatest) + +lemma sup_bot [simp]: + "x \ bot = x" + using bot_least [of x] by (simp add: le_iff_sup sup_commute) + +lemma inf_top [simp]: + "x \ top = x" + using top_greatest [of x] by (simp add: le_iff_inf inf_commute) + +definition SUPR :: "'b set \ ('b \ 'a) \ 'a" where + "SUPR A f == \ (f ` A)" + +definition INFI :: "'b set \ ('b \ 'a) \ 'a" where + "INFI A f == \ (f ` A)" + +end + +syntax + "_SUP1" :: "pttrns => 'b => 'b" ("(3SUP _./ _)" [0, 10] 10) + "_SUP" :: "pttrn => 'a set => 'b => 'b" ("(3SUP _:_./ _)" [0, 10] 10) + "_INF1" :: "pttrns => 'b => 'b" ("(3INF _./ _)" [0, 10] 10) + "_INF" :: "pttrn => 'a set => 'b => 'b" ("(3INF _:_./ _)" [0, 10] 10) + +translations + "SUP x y. B" == "SUP x. SUP y. B" + "SUP x. B" == "CONST SUPR CONST UNIV (%x. B)" + "SUP x. B" == "SUP x:CONST UNIV. B" + "SUP x:A. B" == "CONST SUPR A (%x. B)" + "INF x y. B" == "INF x. INF y. B" + "INF x. B" == "CONST INFI CONST UNIV (%x. B)" + "INF x. B" == "INF x:CONST UNIV. B" + "INF x:A. B" == "CONST INFI A (%x. B)" + +(* To avoid eta-contraction of body: *) +print_translation {* +let + fun btr' syn (A :: Abs abs :: ts) = + let val (x,t) = atomic_abs_tr' abs + in list_comb (Syntax.const syn $ x $ A $ t, ts) end + val const_syntax_name = Sign.const_syntax_name @{theory} o fst o dest_Const +in +[(const_syntax_name @{term SUPR}, btr' "_SUP"),(const_syntax_name @{term "INFI"}, btr' "_INF")] +end +*} + +context complete_lattice +begin + +lemma le_SUPI: "i : A \ M i \ (SUP i:A. M i)" + by (auto simp add: SUPR_def intro: Sup_upper) + +lemma SUP_leI: "(\i. i : A \ M i \ u) \ (SUP i:A. M i) \ u" + by (auto simp add: SUPR_def intro: Sup_least) + +lemma INF_leI: "i : A \ (INF i:A. M i) \ M i" + by (auto simp add: INFI_def intro: Inf_lower) + +lemma le_INFI: "(\i. i : A \ u \ M i) \ u \ (INF i:A. M i)" + by (auto simp add: INFI_def intro: Inf_greatest) + +lemma SUP_const[simp]: "A \ {} \ (SUP i:A. M) = M" + by (auto intro: antisym SUP_leI le_SUPI) + +lemma INF_const[simp]: "A \ {} \ (INF i:A. M) = M" + by (auto intro: antisym INF_leI le_INFI) + +end + + +subsection {* Bool as complete lattice *} + +instantiation bool :: complete_lattice +begin + +definition + Inf_bool_def: "\A \ (\x\A. x)" + +definition + Sup_bool_def: "\A \ (\x\A. x)" + +instance + by intro_classes (auto simp add: Inf_bool_def Sup_bool_def le_bool_def) + +end + +lemma Inf_empty_bool [simp]: + "\{}" + unfolding Inf_bool_def by auto + +lemma not_Sup_empty_bool [simp]: + "\ Sup {}" + unfolding Sup_bool_def by auto + + +subsection {* Fun as complete lattice *} + +instantiation "fun" :: (type, complete_lattice) complete_lattice +begin + +definition + Inf_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" + +definition + Sup_fun_def [code del]: "\A = (\x. \{y. \f\A. y = f x})" + +instance + by intro_classes + (auto simp add: Inf_fun_def Sup_fun_def le_fun_def + intro: Inf_lower Sup_upper Inf_greatest Sup_least) + +end + +lemma Inf_empty_fun: + "\{} = (\_. \{})" + by rule (auto simp add: Inf_fun_def) + +lemma Sup_empty_fun: + "\{} = (\_. \{})" + by rule (auto simp add: Sup_fun_def) + + +subsection {* Set as lattice *} + +lemma inf_set_eq: "A \ B = A \ B" + apply (rule subset_antisym) + apply (rule Int_greatest) + apply (rule inf_le1) + apply (rule inf_le2) + apply (rule inf_greatest) + apply (rule Int_lower1) + apply (rule Int_lower2) + done + +lemma sup_set_eq: "A \ B = A \ B" + apply (rule subset_antisym) + apply (rule sup_least) + apply (rule Un_upper1) + apply (rule Un_upper2) + apply (rule Un_least) + apply (rule sup_ge1) + apply (rule sup_ge2) + done + +lemma mono_Int: "mono f \ f (A \ B) \ f A \ f B" + apply (fold inf_set_eq sup_set_eq) + apply (erule mono_inf) + done + +lemma mono_Un: "mono f \ f A \ f B \ f (A \ B)" + apply (fold inf_set_eq sup_set_eq) + apply (erule mono_sup) + done + +lemma Inf_set_eq: "\S = \S" + apply (rule subset_antisym) + apply (rule Inter_greatest) + apply (erule Inf_lower) + apply (rule Inf_greatest) + apply (erule Inter_lower) + done + +lemma Sup_set_eq: "\S = \S" + apply (rule subset_antisym) + apply (rule Sup_least) + apply (erule Union_upper) + apply (rule Union_least) + apply (erule Sup_upper) + done + +lemma top_set_eq: "top = UNIV" + by (iprover intro!: subset_antisym subset_UNIV top_greatest) + +lemma bot_set_eq: "bot = {}" + by (iprover intro!: subset_antisym empty_subsetI bot_least) + +no_notation + less_eq (infix "\" 50) and + less (infix "\" 50) and + inf (infixl "\" 70) and + sup (infixl "\" 65) and + Inf ("\_" [900] 900) and + Sup ("\_" [900] 900) + subsection {* Basic ML bindings *} diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/SizeChange/sct.ML --- a/src/HOL/SizeChange/sct.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/SizeChange/sct.ML Fri Mar 06 11:10:57 2009 +0100 @@ -266,12 +266,12 @@ fun mk_edge m G n = HOLogic.mk_prod (m, HOLogic.mk_prod (G, n)) -fun mk_set T [] = Const ("{}", HOLogic.mk_setT T) - | mk_set T (x :: xs) = Const ("insert", +fun mk_set T [] = Const (@{const_name Set.empty}, HOLogic.mk_setT T) + | mk_set T (x :: xs) = Const (@{const_name insert}, T --> HOLogic.mk_setT T --> HOLogic.mk_setT T) $ x $ mk_set T xs -fun dest_set (Const ("{}", _)) = [] - | dest_set (Const ("insert", _) $ x $ xs) = x :: dest_set xs +fun dest_set (Const (@{const_name Set.empty}, _)) = [] + | dest_set (Const (@{const_name insert}, _) $ x $ xs) = x :: dest_set xs val in_graph_tac = simp_tac (HOL_basic_ss addsimps has_edge_simps) 1 diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/Qelim/cooper.ML --- a/src/HOL/Tools/Qelim/cooper.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/Qelim/cooper.ML Fri Mar 06 11:10:57 2009 +0100 @@ -428,8 +428,8 @@ in fun provein x S = case term_of S of - Const("{}",_) => error "Unexpected error in Cooper please email Amine Chaieb" - | Const("insert",_)$y$_ => + Const(@{const_name Set.empty}, _) => error "Unexpected error in Cooper please email Amine Chaieb" + | Const(@{const_name insert}, _) $ y $ _ => let val (cy,S') = Thm.dest_binop S in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1 else implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2) diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/Qelim/ferrante_rackoff.ML --- a/src/HOL/Tools/Qelim/ferrante_rackoff.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/Qelim/ferrante_rackoff.ML Fri Mar 06 11:10:57 2009 +0100 @@ -99,8 +99,8 @@ in fun provein x S = case term_of S of - Const("{}",_) => raise CTERM ("provein : not a member!", [S]) - | Const("insert",_)$y$_ => + Const(@{const_name Set.empty}, _) => raise CTERM ("provein : not a member!", [S]) + | Const(@{const_name insert}, _) $ y $_ => let val (cy,S') = Thm.dest_binop S in if term_of x aconv y then instantiate' [] [SOME x, SOME S'] insI1 else implies_elim (instantiate' [] [SOME x, SOME S', SOME cy] insI2) diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/Qelim/langford.ML --- a/src/HOL/Tools/Qelim/langford.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/Qelim/langford.ML Fri Mar 06 11:10:57 2009 +0100 @@ -14,9 +14,9 @@ val dest_set = let fun h acc ct = - case (term_of ct) of - Const("{}",_) => acc - | Const("insert",_)$_$t => h (Thm.dest_arg1 ct :: acc) (Thm.dest_arg ct) + case term_of ct of + Const (@{const_name Set.empty}, _) => acc + | Const (@{const_name insert}, _) $ _ $ t => h (Thm.dest_arg1 ct :: acc) (Thm.dest_arg ct) in h [] end; fun prove_finite cT u = @@ -48,11 +48,11 @@ in (ne, f) end val qe = case (term_of L, term_of U) of - (Const("{}",_),_) => + (Const (@{const_name Set.empty}, _),_) => let val (neU,fU) = proveneF U in simp_rule (transitive ths (dlo_qeth_nolb OF [neU, fU])) end - | (_,Const("{}",_)) => + | (_,Const (@{const_name Set.empty}, _)) => let val (neL,fL) = proveneF L in simp_rule (transitive ths (dlo_qeth_noub OF [neL, fL])) end diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/datatype_rep_proofs.ML --- a/src/HOL/Tools/datatype_rep_proofs.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/datatype_rep_proofs.ML Fri Mar 06 11:10:57 2009 +0100 @@ -461,7 +461,7 @@ (if i < length newTs then Const ("True", HOLogic.boolT) else HOLogic.mk_mem (mk_Free "x" Univ_elT i, Const ("image", [isoT, HOLogic.mk_setT T] ---> UnivT) $ - Const (iso_name, isoT) $ Const ("UNIV", HOLogic.mk_setT T))) + Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T))) end; val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/function_package/decompose.ML --- a/src/HOL/Tools/function_package/decompose.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/function_package/decompose.ML Fri Mar 06 11:10:57 2009 +0100 @@ -30,7 +30,7 @@ if is_some (Termination.get_chain D c1 c2) then D else let val goal = HOLogic.mk_eq (HOLogic.mk_binop @{const_name "Relation.rel_comp"} (c1, c2), - Const (@{const_name "{}"}, fastype_of c1)) + Const (@{const_name Set.empty}, fastype_of c1)) |> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *) val chain = case FundefLib.try_proof (cterm_of thy goal) chain_tac of diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/function_package/fundef_lib.ML --- a/src/HOL/Tools/function_package/fundef_lib.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/function_package/fundef_lib.ML Fri Mar 06 11:10:57 2009 +0100 @@ -139,7 +139,7 @@ Here, + can be any binary operation that is AC. - cn - The name of the binop-constructor (e.g. @{const_name "op Un"}) + cn - The name of the binop-constructor (e.g. @{const_name Un}) ac - the AC rewrite rules for cn is - the list of indices of the expressions that should become the first part (e.g. [0,1,3] in the above example) @@ -168,8 +168,8 @@ (* instance for unions *) fun regroup_union_conv t = - regroup_conv (@{const_name "{}"}) - @{const_name "op Un"} + regroup_conv (@{const_name Set.empty}) + @{const_name Un} (map (fn t => t RS eq_reflection) (@{thms "Un_ac"} @ @{thms "Un_empty_right"} @ @{thms "Un_empty_left"})) t diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/function_package/lexicographic_order.ML --- a/src/HOL/Tools/function_package/lexicographic_order.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/function_package/lexicographic_order.ML Fri Mar 06 11:10:57 2009 +0100 @@ -24,7 +24,7 @@ let val relT = HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) val mlexT = (domT --> HOLogic.natT) --> relT --> relT - fun mk_ms [] = Const (@{const_name "{}"}, relT) + fun mk_ms [] = Const (@{const_name Set.empty}, relT) | mk_ms (f::fs) = Const (@{const_name "mlex_prod"}, mlexT) $ f $ mk_ms fs in diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/function_package/scnp_reconstruct.ML --- a/src/HOL/Tools/function_package/scnp_reconstruct.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/function_package/scnp_reconstruct.ML Fri Mar 06 11:10:57 2009 +0100 @@ -142,7 +142,7 @@ val setT = HOLogic.mk_setT -fun mk_set T [] = Const (@{const_name "{}"}, setT T) +fun mk_set T [] = Const (@{const_name Set.empty}, setT T) | mk_set T (x :: xs) = Const (@{const_name insert}, T --> setT T --> setT T) $ x $ mk_set T xs diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/function_package/termination.ML --- a/src/HOL/Tools/function_package/termination.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/function_package/termination.ML Fri Mar 06 11:10:57 2009 +0100 @@ -145,7 +145,7 @@ fun mk_sum_skel rel = let - val cs = FundefLib.dest_binop_list @{const_name "op Un"} rel + val cs = FundefLib.dest_binop_list @{const_name Un} rel fun collect_pats (Const ("Collect", _) $ Abs (_, _, c)) = let val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam) @@ -233,7 +233,7 @@ fun CALLS tac i st = if Thm.no_prems st then all_tac st else case Thm.term_of (Thm.cprem_of st i) of - (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name "op Un"} rel, i) st + (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name Un} rel, i) st |_ => no_tac st type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic @@ -291,9 +291,9 @@ val relation = if null ineqs then - Const (@{const_name "{}"}, fastype_of rel) + Const (@{const_name Set.empty}, fastype_of rel) else - foldr1 (HOLogic.mk_binop @{const_name "op Un"}) (map mk_compr ineqs) + foldr1 (HOLogic.mk_binop @{const_name Un}) (map mk_compr ineqs) fun solve_membership_tac i = (EVERY' (replicate (i - 2) (rtac @{thm UnI2})) (* pick the right component of the union *) diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/hologic.ML --- a/src/HOL/Tools/hologic.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/hologic.ML Fri Mar 06 11:10:57 2009 +0100 @@ -225,7 +225,7 @@ fun dest_mem (Const ("op :", _) $ x $ A) = (x, A) | dest_mem t = raise TERM ("dest_mem", [t]); -fun mk_UNIV T = Const ("UNIV", mk_setT T); +fun mk_UNIV T = Const ("Set.UNIV", mk_setT T); (* binary operations and relations *) diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/inductive_set_package.ML --- a/src/HOL/Tools/inductive_set_package.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/inductive_set_package.ML Fri Mar 06 11:10:57 2009 +0100 @@ -73,8 +73,8 @@ in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs)) (p (fold (Logic.all o Var) vs t) f) end; - fun mkop "op &" T x = SOME (Const ("op Int", T --> T --> T), x) - | mkop "op |" T x = SOME (Const ("op Un", T --> T --> T), x) + fun mkop "op &" T x = SOME (Const (@{const_name "Int"}, T --> T --> T), x) + | mkop "op |" T x = SOME (Const (@{const_name "Un"}, T --> T --> T), x) | mkop _ _ _ = NONE; fun mk_collect p T t = let val U = HOLogic.dest_setT T diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/refute.ML --- a/src/HOL/Tools/refute.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/refute.ML Fri Mar 06 11:10:57 2009 +0100 @@ -2094,7 +2094,7 @@ val HOLogic_prodT = HOLogic.mk_prodT (T1, T2) val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT (* Term.term *) - val HOLogic_empty_set = Const (@{const_name "{}"}, HOLogic_setT) + val HOLogic_empty_set = Const (@{const_name Set.empty}, HOLogic_setT) val HOLogic_insert = Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT) in @@ -3159,7 +3159,7 @@ val HOLogic_prodT = HOLogic.mk_prodT (T1, T2) val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT (* Term.term *) - val HOLogic_empty_set = Const (@{const_name "{}"}, HOLogic_setT) + val HOLogic_empty_set = Const (@{const_name Set.empty}, HOLogic_setT) val HOLogic_insert = Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT) in diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/Tools/res_clause.ML --- a/src/HOL/Tools/res_clause.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/Tools/res_clause.ML Fri Mar 06 11:10:57 2009 +0100 @@ -98,21 +98,21 @@ (*Provide readable names for the more common symbolic functions*) val const_trans_table = - Symtab.make [("op =", "equal"), + Symtab.make [(@{const_name "op ="}, "equal"), (@{const_name HOL.less_eq}, "lessequals"), (@{const_name HOL.less}, "less"), - ("op &", "and"), - ("op |", "or"), + (@{const_name "op &"}, "and"), + (@{const_name "op |"}, "or"), (@{const_name HOL.plus}, "plus"), (@{const_name HOL.minus}, "minus"), (@{const_name HOL.times}, "times"), (@{const_name Divides.div}, "div"), (@{const_name HOL.divide}, "divide"), - ("op -->", "implies"), - ("{}", "emptyset"), - ("op :", "in"), - ("op Un", "union"), - ("op Int", "inter"), + (@{const_name "op -->"}, "implies"), + (@{const_name Set.empty}, "emptyset"), + (@{const_name "op :"}, "in"), + (@{const_name Un}, "union"), + (@{const_name Int}, "inter"), ("List.append", "append"), ("ATP_Linkup.fequal", "fequal"), ("ATP_Linkup.COMBI", "COMBI"), diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOL/UNITY/Union.thy --- a/src/HOL/UNITY/Union.thy Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOL/UNITY/Union.thy Fri Mar 06 11:10:57 2009 +0100 @@ -43,7 +43,7 @@ translations "JN x : A. B" == "JOIN A (%x. B)" "JN x y. B" == "JN x. JN y. B" - "JN x. B" == "JOIN UNIV (%x. B)" + "JN x. B" == "JOIN CONST UNIV (%x. B)" syntax (xsymbols) SKIP :: "'a program" ("\") diff -r d6bffd97d8d5 -r 6c74ef5a349f src/HOLCF/IOA/meta_theory/ioa_package.ML --- a/src/HOLCF/IOA/meta_theory/ioa_package.ML Fri Mar 06 09:35:43 2009 +0100 +++ b/src/HOLCF/IOA/meta_theory/ioa_package.ML Fri Mar 06 11:10:57 2009 +0100 @@ -159,7 +159,8 @@ (constructor_head thy atypstr a)::(constructor_head_list thy atypstr r); (* producing an action set *) -fun action_set_string thy atyp [] = "{}" | +(*FIXME*) +fun action_set_string thy atyp [] = "Set.empty" | action_set_string thy atyp [a] = constr_set_string thy atyp (strip_quote a) | action_set_string thy atyp (a::r) = (constr_set_string thy atyp (strip_quote a)) ^ " Un " ^ (action_set_string thy atyp r);