src/HOL/Orderings.thy
author wenzelm
Tue, 09 Oct 2007 00:20:13 +0200
changeset 24920 2a45e400fdad
parent 24867 e5b55d7be9bb
child 25062 af5ef0d4d655
permissions -rw-r--r--
generic Syntax.pretty/string_of operations;
     1 (*  Title:      HOL/Orderings.thy
     2     ID:         $Id$
     3     Author:     Tobias Nipkow, Markus Wenzel, and Larry Paulson
     4 *)
     5 
     6 header {* Syntactic and abstract orders *}
     7 
     8 theory Orderings
     9 imports Set Fun
    10 uses
    11   "~~/src/Provers/order.ML"
    12 begin
    13 
    14 subsection {* Partial orders *}
    15 
    16 class order = ord +
    17   assumes less_le: "x \<^loc>< y \<longleftrightarrow> x \<^loc>\<le> y \<and> x \<noteq> y"
    18   and order_refl [iff]: "x \<^loc>\<le> x"
    19   and order_trans: "x \<^loc>\<le> y \<Longrightarrow> y \<^loc>\<le> z \<Longrightarrow> x \<^loc>\<le> z"
    20   assumes antisym: "x \<^loc>\<le> y \<Longrightarrow> y \<^loc>\<le> x \<Longrightarrow> x = y"
    21 
    22 begin
    23 
    24 notation (input)
    25   less_eq (infix "\<sqsubseteq>" 50)
    26 and
    27   less    (infix "\<sqsubset>" 50)
    28 
    29 text {* Reflexivity. *}
    30 
    31 lemma eq_refl: "x = y \<Longrightarrow> x \<^loc>\<le> y"
    32     -- {* This form is useful with the classical reasoner. *}
    33 by (erule ssubst) (rule order_refl)
    34 
    35 lemma less_irrefl [iff]: "\<not> x \<^loc>< x"
    36 by (simp add: less_le)
    37 
    38 lemma le_less: "x \<^loc>\<le> y \<longleftrightarrow> x \<^loc>< y \<or> x = y"
    39     -- {* NOT suitable for iff, since it can cause PROOF FAILED. *}
    40 by (simp add: less_le) blast
    41 
    42 lemma le_imp_less_or_eq: "x \<^loc>\<le> y \<Longrightarrow> x \<^loc>< y \<or> x = y"
    43 unfolding less_le by blast
    44 
    45 lemma less_imp_le: "x \<^loc>< y \<Longrightarrow> x \<^loc>\<le> y"
    46 unfolding less_le by blast
    47 
    48 lemma less_imp_neq: "x \<^loc>< y \<Longrightarrow> x \<noteq> y"
    49 by (erule contrapos_pn, erule subst, rule less_irrefl)
    50 
    51 
    52 text {* Useful for simplification, but too risky to include by default. *}
    53 
    54 lemma less_imp_not_eq: "x \<^loc>< y \<Longrightarrow> (x = y) \<longleftrightarrow> False"
    55 by auto
    56 
    57 lemma less_imp_not_eq2: "x \<^loc>< y \<Longrightarrow> (y = x) \<longleftrightarrow> False"
    58 by auto
    59 
    60 
    61 text {* Transitivity rules for calculational reasoning *}
    62 
    63 lemma neq_le_trans: "a \<noteq> b \<Longrightarrow> a \<^loc>\<le> b \<Longrightarrow> a \<^loc>< b"
    64 by (simp add: less_le)
    65 
    66 lemma le_neq_trans: "a \<^loc>\<le> b \<Longrightarrow> a \<noteq> b \<Longrightarrow> a \<^loc>< b"
    67 by (simp add: less_le)
    68 
    69 
    70 text {* Asymmetry. *}
    71 
    72 lemma less_not_sym: "x \<^loc>< y \<Longrightarrow> \<not> (y \<^loc>< x)"
    73 by (simp add: less_le antisym)
    74 
    75 lemma less_asym: "x \<^loc>< y \<Longrightarrow> (\<not> P \<Longrightarrow> y \<^loc>< x) \<Longrightarrow> P"
    76 by (drule less_not_sym, erule contrapos_np) simp
    77 
    78 lemma eq_iff: "x = y \<longleftrightarrow> x \<^loc>\<le> y \<and> y \<^loc>\<le> x"
    79 by (blast intro: antisym)
    80 
    81 lemma antisym_conv: "y \<^loc>\<le> x \<Longrightarrow> x \<^loc>\<le> y \<longleftrightarrow> x = y"
    82 by (blast intro: antisym)
    83 
    84 lemma less_imp_neq: "x \<^loc>< y \<Longrightarrow> x \<noteq> y"
    85 by (erule contrapos_pn, erule subst, rule less_irrefl)
    86 
    87 
    88 text {* Transitivity. *}
    89 
    90 lemma less_trans: "x \<^loc>< y \<Longrightarrow> y \<^loc>< z \<Longrightarrow> x \<^loc>< z"
    91 by (simp add: less_le) (blast intro: order_trans antisym)
    92 
    93 lemma le_less_trans: "x \<^loc>\<le> y \<Longrightarrow> y \<^loc>< z \<Longrightarrow> x \<^loc>< z"
    94 by (simp add: less_le) (blast intro: order_trans antisym)
    95 
    96 lemma less_le_trans: "x \<^loc>< y \<Longrightarrow> y \<^loc>\<le> z \<Longrightarrow> x \<^loc>< z"
    97 by (simp add: less_le) (blast intro: order_trans antisym)
    98 
    99 
   100 text {* Useful for simplification, but too risky to include by default. *}
   101 
   102 lemma less_imp_not_less: "x \<^loc>< y \<Longrightarrow> (\<not> y \<^loc>< x) \<longleftrightarrow> True"
   103 by (blast elim: less_asym)
   104 
   105 lemma less_imp_triv: "x \<^loc>< y \<Longrightarrow> (y \<^loc>< x \<longrightarrow> P) \<longleftrightarrow> True"
   106 by (blast elim: less_asym)
   107 
   108 
   109 text {* Transitivity rules for calculational reasoning *}
   110 
   111 lemma less_asym': "a \<^loc>< b \<Longrightarrow> b \<^loc>< a \<Longrightarrow> P"
   112 by (rule less_asym)
   113 
   114 
   115 text {* Reverse order *}
   116 
   117 lemma order_reverse:
   118   "order (\<lambda>x y. y \<^loc>\<le> x) (\<lambda>x y. y \<^loc>< x)"
   119 by unfold_locales
   120    (simp add: less_le, auto intro: antisym order_trans)
   121 
   122 end
   123 
   124 
   125 subsection {* Linear (total) orders *}
   126 
   127 class linorder = order +
   128   assumes linear: "x \<sqsubseteq> y \<or> y \<sqsubseteq> x"
   129 begin
   130 
   131 lemma less_linear: "x \<^loc>< y \<or> x = y \<or> y \<^loc>< x"
   132 unfolding less_le using less_le linear by blast
   133 
   134 lemma le_less_linear: "x \<^loc>\<le> y \<or> y \<^loc>< x"
   135 by (simp add: le_less less_linear)
   136 
   137 lemma le_cases [case_names le ge]:
   138   "(x \<^loc>\<le> y \<Longrightarrow> P) \<Longrightarrow> (y \<^loc>\<le> x \<Longrightarrow> P) \<Longrightarrow> P"
   139 using linear by blast
   140 
   141 lemma linorder_cases [case_names less equal greater]:
   142   "(x \<^loc>< y \<Longrightarrow> P) \<Longrightarrow> (x = y \<Longrightarrow> P) \<Longrightarrow> (y \<^loc>< x \<Longrightarrow> P) \<Longrightarrow> P"
   143 using less_linear by blast
   144 
   145 lemma not_less: "\<not> x \<^loc>< y \<longleftrightarrow> y \<^loc>\<le> x"
   146 apply (simp add: less_le)
   147 using linear apply (blast intro: antisym)
   148 done
   149 
   150 lemma not_less_iff_gr_or_eq:
   151  "\<not>(x \<^loc>< y) \<longleftrightarrow> (x \<^loc>> y | x = y)"
   152 apply(simp add:not_less le_less)
   153 apply blast
   154 done
   155 
   156 lemma not_le: "\<not> x \<^loc>\<le> y \<longleftrightarrow> y \<^loc>< x"
   157 apply (simp add: less_le)
   158 using linear apply (blast intro: antisym)
   159 done
   160 
   161 lemma neq_iff: "x \<noteq> y \<longleftrightarrow> x \<^loc>< y \<or> y \<^loc>< x"
   162 by (cut_tac x = x and y = y in less_linear, auto)
   163 
   164 lemma neqE: "x \<noteq> y \<Longrightarrow> (x \<^loc>< y \<Longrightarrow> R) \<Longrightarrow> (y \<^loc>< x \<Longrightarrow> R) \<Longrightarrow> R"
   165 by (simp add: neq_iff) blast
   166 
   167 lemma antisym_conv1: "\<not> x \<^loc>< y \<Longrightarrow> x \<^loc>\<le> y \<longleftrightarrow> x = y"
   168 by (blast intro: antisym dest: not_less [THEN iffD1])
   169 
   170 lemma antisym_conv2: "x \<^loc>\<le> y \<Longrightarrow> \<not> x \<^loc>< y \<longleftrightarrow> x = y"
   171 by (blast intro: antisym dest: not_less [THEN iffD1])
   172 
   173 lemma antisym_conv3: "\<not> y \<^loc>< x \<Longrightarrow> \<not> x \<^loc>< y \<longleftrightarrow> x = y"
   174 by (blast intro: antisym dest: not_less [THEN iffD1])
   175 
   176 text{*Replacing the old Nat.leI*}
   177 lemma leI: "\<not> x \<^loc>< y \<Longrightarrow> y \<^loc>\<le> x"
   178 unfolding not_less .
   179 
   180 lemma leD: "y \<^loc>\<le> x \<Longrightarrow> \<not> x \<^loc>< y"
   181 unfolding not_less .
   182 
   183 (*FIXME inappropriate name (or delete altogether)*)
   184 lemma not_leE: "\<not> y \<^loc>\<le> x \<Longrightarrow> x \<^loc>< y"
   185 unfolding not_le .
   186 
   187 
   188 text {* Reverse order *}
   189 
   190 lemma linorder_reverse:
   191   "linorder (\<lambda>x y. y \<^loc>\<le> x) (\<lambda>x y. y \<^loc>< x)"
   192 by unfold_locales
   193   (simp add: less_le, auto intro: antisym order_trans simp add: linear)
   194 
   195 
   196 text {* min/max *}
   197 
   198 text {* for historic reasons, definitions are done in context ord *}
   199 
   200 definition (in ord)
   201   min :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
   202   [code unfold, code inline del]: "min a b = (if a \<^loc>\<le> b then a else b)"
   203 
   204 definition (in ord)
   205   max :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" where
   206   [code unfold, code inline del]: "max a b = (if a \<^loc>\<le> b then b else a)"
   207 
   208 lemma min_le_iff_disj:
   209   "min x y \<^loc>\<le> z \<longleftrightarrow> x \<^loc>\<le> z \<or> y \<^loc>\<le> z"
   210 unfolding min_def using linear by (auto intro: order_trans)
   211 
   212 lemma le_max_iff_disj:
   213   "z \<^loc>\<le> max x y \<longleftrightarrow> z \<^loc>\<le> x \<or> z \<^loc>\<le> y"
   214 unfolding max_def using linear by (auto intro: order_trans)
   215 
   216 lemma min_less_iff_disj:
   217   "min x y \<^loc>< z \<longleftrightarrow> x \<^loc>< z \<or> y \<^loc>< z"
   218 unfolding min_def le_less using less_linear by (auto intro: less_trans)
   219 
   220 lemma less_max_iff_disj:
   221   "z \<^loc>< max x y \<longleftrightarrow> z \<^loc>< x \<or> z \<^loc>< y"
   222 unfolding max_def le_less using less_linear by (auto intro: less_trans)
   223 
   224 lemma min_less_iff_conj [simp]:
   225   "z \<^loc>< min x y \<longleftrightarrow> z \<^loc>< x \<and> z \<^loc>< y"
   226 unfolding min_def le_less using less_linear by (auto intro: less_trans)
   227 
   228 lemma max_less_iff_conj [simp]:
   229   "max x y \<^loc>< z \<longleftrightarrow> x \<^loc>< z \<and> y \<^loc>< z"
   230 unfolding max_def le_less using less_linear by (auto intro: less_trans)
   231 
   232 lemma split_min [noatp]:
   233   "P (min i j) \<longleftrightarrow> (i \<^loc>\<le> j \<longrightarrow> P i) \<and> (\<not> i \<^loc>\<le> j \<longrightarrow> P j)"
   234 by (simp add: min_def)
   235 
   236 lemma split_max [noatp]:
   237   "P (max i j) \<longleftrightarrow> (i \<^loc>\<le> j \<longrightarrow> P j) \<and> (\<not> i \<^loc>\<le> j \<longrightarrow> P i)"
   238 by (simp add: max_def)
   239 
   240 end
   241 
   242 
   243 subsection {* Reasoning tools setup *}
   244 
   245 ML {*
   246 
   247 signature ORDERS =
   248 sig
   249   val print_structures: Proof.context -> unit
   250   val setup: theory -> theory
   251   val order_tac: thm list -> Proof.context -> int -> tactic
   252 end;
   253 
   254 structure Orders: ORDERS =
   255 struct
   256 
   257 (** Theory and context data **)
   258 
   259 fun struct_eq ((s1: string, ts1), (s2, ts2)) =
   260   (s1 = s2) andalso eq_list (op aconv) (ts1, ts2);
   261 
   262 structure Data = GenericDataFun
   263 (
   264   type T = ((string * term list) * Order_Tac.less_arith) list;
   265     (* Order structures:
   266        identifier of the structure, list of operations and record of theorems
   267        needed to set up the transitivity reasoner,
   268        identifier and operations identify the structure uniquely. *)
   269   val empty = [];
   270   val extend = I;
   271   fun merge _ = AList.join struct_eq (K fst);
   272 );
   273 
   274 fun print_structures ctxt =
   275   let
   276     val structs = Data.get (Context.Proof ctxt);
   277     fun pretty_term t = Pretty.block
   278       [Pretty.quote (Syntax.pretty_term ctxt t), Pretty.brk 1,
   279         Pretty.str "::", Pretty.brk 1,
   280         Pretty.quote (Syntax.pretty_typ ctxt (type_of t))];
   281     fun pretty_struct ((s, ts), _) = Pretty.block
   282       [Pretty.str s, Pretty.str ":", Pretty.brk 1,
   283        Pretty.enclose "(" ")" (Pretty.breaks (map pretty_term ts))];
   284   in
   285     Pretty.writeln (Pretty.big_list "Order structures:" (map pretty_struct structs))
   286   end;
   287 
   288 
   289 (** Method **)
   290 
   291 fun struct_tac ((s, [eq, le, less]), thms) prems =
   292   let
   293     fun decomp thy (Trueprop $ t) =
   294       let
   295         fun excluded t =
   296           (* exclude numeric types: linear arithmetic subsumes transitivity *)
   297           let val T = type_of t
   298           in
   299 	    T = HOLogic.natT orelse T = HOLogic.intT orelse T = HOLogic.realT
   300           end;
   301 	fun rel (bin_op $ t1 $ t2) =
   302               if excluded t1 then NONE
   303               else if Pattern.matches thy (eq, bin_op) then SOME (t1, "=", t2)
   304               else if Pattern.matches thy (le, bin_op) then SOME (t1, "<=", t2)
   305               else if Pattern.matches thy (less, bin_op) then SOME (t1, "<", t2)
   306               else NONE
   307 	  | rel _ = NONE;
   308 	fun dec (Const (@{const_name Not}, _) $ t) = (case rel t
   309 	      of NONE => NONE
   310 	       | SOME (t1, rel, t2) => SOME (t1, "~" ^ rel, t2))
   311           | dec x = rel x;
   312       in dec t end;
   313   in
   314     case s of
   315       "order" => Order_Tac.partial_tac decomp thms prems
   316     | "linorder" => Order_Tac.linear_tac decomp thms prems
   317     | _ => error ("Unknown kind of order `" ^ s ^ "' encountered in transitivity reasoner.")
   318   end
   319 
   320 fun order_tac prems ctxt =
   321   FIRST' (map (fn s => CHANGED o struct_tac s prems) (Data.get (Context.Proof ctxt)));
   322 
   323 
   324 (** Attribute **)
   325 
   326 fun add_struct_thm s tag =
   327   Thm.declaration_attribute
   328     (fn thm => Data.map (AList.map_default struct_eq (s, Order_Tac.empty TrueI) (Order_Tac.update tag thm)));
   329 fun del_struct s =
   330   Thm.declaration_attribute
   331     (fn _ => Data.map (AList.delete struct_eq s));
   332 
   333 val attribute = Attrib.syntax
   334      (Scan.lift ((Args.add -- Args.name >> (fn (_, s) => SOME s) ||
   335           Args.del >> K NONE) --| Args.colon (* FIXME ||
   336         Scan.succeed true *) ) -- Scan.lift Args.name --
   337       Scan.repeat Args.term
   338       >> (fn ((SOME tag, n), ts) => add_struct_thm (n, ts) tag
   339            | ((NONE, n), ts) => del_struct (n, ts)));
   340 
   341 
   342 (** Diagnostic command **)
   343 
   344 val print = Toplevel.unknown_context o
   345   Toplevel.keep (Toplevel.node_case
   346     (Context.cases (print_structures o ProofContext.init) print_structures)
   347     (print_structures o Proof.context_of));
   348 
   349 val _ =
   350   OuterSyntax.improper_command "print_orders"
   351     "print order structures available to transitivity reasoner" OuterKeyword.diag
   352     (Scan.succeed (Toplevel.no_timing o print));
   353 
   354 
   355 (** Setup **)
   356 
   357 val setup =
   358   Method.add_methods
   359     [("order", Method.ctxt_args (Method.SIMPLE_METHOD' o order_tac []), "transitivity reasoner")] #>
   360   Attrib.add_attributes [("order", attribute, "theorems controlling transitivity reasoner")];
   361 
   362 end;
   363 
   364 *}
   365 
   366 setup Orders.setup
   367 
   368 
   369 text {* Declarations to set up transitivity reasoner of partial and linear orders. *}
   370 
   371 (* The type constraint on @{term op =} below is necessary since the operation
   372    is not a parameter of the locale. *)
   373 lemmas (in order)
   374   [order add less_reflE: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   375   less_irrefl [THEN notE]
   376 lemmas (in order)
   377   [order add le_refl: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   378   order_refl
   379 lemmas (in order)
   380   [order add less_imp_le: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   381   less_imp_le
   382 lemmas (in order)
   383   [order add eqI: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   384   antisym
   385 lemmas (in order)
   386   [order add eqD1: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   387   eq_refl
   388 lemmas (in order)
   389   [order add eqD2: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   390   sym [THEN eq_refl]
   391 lemmas (in order)
   392   [order add less_trans: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   393   less_trans
   394 lemmas (in order)
   395   [order add less_le_trans: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   396   less_le_trans
   397 lemmas (in order)
   398   [order add le_less_trans: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   399   le_less_trans
   400 lemmas (in order)
   401   [order add le_trans: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   402   order_trans
   403 lemmas (in order)
   404   [order add le_neq_trans: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   405   le_neq_trans
   406 lemmas (in order)
   407   [order add neq_le_trans: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   408   neq_le_trans
   409 lemmas (in order)
   410   [order add less_imp_neq: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   411   less_imp_neq
   412 lemmas (in order)
   413   [order add eq_neq_eq_imp_neq: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   414    eq_neq_eq_imp_neq
   415 lemmas (in order)
   416   [order add not_sym: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   417   not_sym
   418 
   419 lemmas (in linorder) [order del: order "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] = _
   420 
   421 lemmas (in linorder)
   422   [order add less_reflE: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   423   less_irrefl [THEN notE]
   424 lemmas (in linorder)
   425   [order add le_refl: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   426   order_refl
   427 lemmas (in linorder)
   428   [order add less_imp_le: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   429   less_imp_le
   430 lemmas (in linorder)
   431   [order add not_lessI: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   432   not_less [THEN iffD2]
   433 lemmas (in linorder)
   434   [order add not_leI: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   435   not_le [THEN iffD2]
   436 lemmas (in linorder)
   437   [order add not_lessD: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   438   not_less [THEN iffD1]
   439 lemmas (in linorder)
   440   [order add not_leD: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   441   not_le [THEN iffD1]
   442 lemmas (in linorder)
   443   [order add eqI: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   444   antisym
   445 lemmas (in linorder)
   446   [order add eqD1: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   447   eq_refl
   448 lemmas (in linorder)
   449   [order add eqD2: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   450   sym [THEN eq_refl]
   451 lemmas (in linorder)
   452   [order add less_trans: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   453   less_trans
   454 lemmas (in linorder)
   455   [order add less_le_trans: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   456   less_le_trans
   457 lemmas (in linorder)
   458   [order add le_less_trans: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   459   le_less_trans
   460 lemmas (in linorder)
   461   [order add le_trans: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   462   order_trans
   463 lemmas (in linorder)
   464   [order add le_neq_trans: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   465   le_neq_trans
   466 lemmas (in linorder)
   467   [order add neq_le_trans: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   468   neq_le_trans
   469 lemmas (in linorder)
   470   [order add less_imp_neq: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   471   less_imp_neq
   472 lemmas (in linorder)
   473   [order add eq_neq_eq_imp_neq: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   474   eq_neq_eq_imp_neq
   475 lemmas (in linorder)
   476   [order add not_sym: linorder "op = :: 'a => 'a => bool" "op \<^loc><=" "op \<^loc><"] =
   477   not_sym
   478 
   479 
   480 setup {*
   481 let
   482 
   483 fun prp t thm = (#prop (rep_thm thm) = t);
   484 
   485 fun prove_antisym_le sg ss ((le as Const(_,T)) $ r $ s) =
   486   let val prems = prems_of_ss ss;
   487       val less = Const (@{const_name less}, T);
   488       val t = HOLogic.mk_Trueprop(le $ s $ r);
   489   in case find_first (prp t) prems of
   490        NONE =>
   491          let val t = HOLogic.mk_Trueprop(HOLogic.Not $ (less $ r $ s))
   492          in case find_first (prp t) prems of
   493               NONE => NONE
   494             | SOME thm => SOME(mk_meta_eq(thm RS @{thm linorder_class.antisym_conv1}))
   495          end
   496      | SOME thm => SOME(mk_meta_eq(thm RS @{thm order_class.antisym_conv}))
   497   end
   498   handle THM _ => NONE;
   499 
   500 fun prove_antisym_less sg ss (NotC $ ((less as Const(_,T)) $ r $ s)) =
   501   let val prems = prems_of_ss ss;
   502       val le = Const (@{const_name less_eq}, T);
   503       val t = HOLogic.mk_Trueprop(le $ r $ s);
   504   in case find_first (prp t) prems of
   505        NONE =>
   506          let val t = HOLogic.mk_Trueprop(NotC $ (less $ s $ r))
   507          in case find_first (prp t) prems of
   508               NONE => NONE
   509             | SOME thm => SOME(mk_meta_eq(thm RS @{thm linorder_class.antisym_conv3}))
   510          end
   511      | SOME thm => SOME(mk_meta_eq(thm RS @{thm linorder_class.antisym_conv2}))
   512   end
   513   handle THM _ => NONE;
   514 
   515 fun add_simprocs procs thy =
   516   (Simplifier.change_simpset_of thy (fn ss => ss
   517     addsimprocs (map (fn (name, raw_ts, proc) =>
   518       Simplifier.simproc thy name raw_ts proc)) procs); thy);
   519 fun add_solver name tac thy =
   520   (Simplifier.change_simpset_of thy (fn ss => ss addSolver
   521     (mk_solver' name (fn ss => tac (MetaSimplifier.prems_of_ss ss) (MetaSimplifier.the_context ss)))); thy);
   522 
   523 in
   524   add_simprocs [
   525        ("antisym le", ["(x::'a::order) <= y"], prove_antisym_le),
   526        ("antisym less", ["~ (x::'a::linorder) < y"], prove_antisym_less)
   527      ]
   528   #> add_solver "Transitivity" Orders.order_tac
   529   (* Adding the transitivity reasoners also as safe solvers showed a slight
   530      speed up, but the reasoning strength appears to be not higher (at least
   531      no breaking of additional proofs in the entire HOL distribution, as
   532      of 5 March 2004, was observed). *)
   533 end
   534 *}
   535 
   536 
   537 subsection {* Dense orders *}
   538 
   539 class dense_linear_order = linorder + 
   540   assumes gt_ex: "\<exists>y. x \<sqsubset> y" 
   541   and lt_ex: "\<exists>y. y \<sqsubset> x"
   542   and dense: "x \<sqsubset> y \<Longrightarrow> (\<exists>z. x \<sqsubset> z \<and> z \<sqsubset> y)"
   543   (*see further theory Dense_Linear_Order*)
   544 
   545 
   546 lemma interval_empty_iff:
   547   fixes x y z :: "'a\<Colon>dense_linear_order"
   548   shows "{y. x < y \<and> y < z} = {} \<longleftrightarrow> \<not> x < z"
   549   by (auto dest: dense)
   550 
   551 subsection {* Name duplicates *}
   552 
   553 lemmas order_less_le = less_le
   554 lemmas order_eq_refl = order_class.eq_refl
   555 lemmas order_less_irrefl = order_class.less_irrefl
   556 lemmas order_le_less = order_class.le_less
   557 lemmas order_le_imp_less_or_eq = order_class.le_imp_less_or_eq
   558 lemmas order_less_imp_le = order_class.less_imp_le
   559 lemmas order_less_imp_not_eq = order_class.less_imp_not_eq
   560 lemmas order_less_imp_not_eq2 = order_class.less_imp_not_eq2
   561 lemmas order_neq_le_trans = order_class.neq_le_trans
   562 lemmas order_le_neq_trans = order_class.le_neq_trans
   563 
   564 lemmas order_antisym = antisym
   565 lemmas order_less_not_sym = order_class.less_not_sym
   566 lemmas order_less_asym = order_class.less_asym
   567 lemmas order_eq_iff = order_class.eq_iff
   568 lemmas order_antisym_conv = order_class.antisym_conv
   569 lemmas order_less_trans = order_class.less_trans
   570 lemmas order_le_less_trans = order_class.le_less_trans
   571 lemmas order_less_le_trans = order_class.less_le_trans
   572 lemmas order_less_imp_not_less = order_class.less_imp_not_less
   573 lemmas order_less_imp_triv = order_class.less_imp_triv
   574 lemmas order_less_asym' = order_class.less_asym'
   575 
   576 lemmas linorder_linear = linear
   577 lemmas linorder_less_linear = linorder_class.less_linear
   578 lemmas linorder_le_less_linear = linorder_class.le_less_linear
   579 lemmas linorder_le_cases = linorder_class.le_cases
   580 lemmas linorder_not_less = linorder_class.not_less
   581 lemmas linorder_not_le = linorder_class.not_le
   582 lemmas linorder_neq_iff = linorder_class.neq_iff
   583 lemmas linorder_neqE = linorder_class.neqE
   584 lemmas linorder_antisym_conv1 = linorder_class.antisym_conv1
   585 lemmas linorder_antisym_conv2 = linorder_class.antisym_conv2
   586 lemmas linorder_antisym_conv3 = linorder_class.antisym_conv3
   587 
   588 lemmas min_le_iff_disj = linorder_class.min_le_iff_disj
   589 lemmas le_max_iff_disj = linorder_class.le_max_iff_disj
   590 lemmas min_less_iff_disj = linorder_class.min_less_iff_disj
   591 lemmas less_max_iff_disj = linorder_class.less_max_iff_disj
   592 lemmas min_less_iff_conj [simp] = linorder_class.min_less_iff_conj
   593 lemmas max_less_iff_conj [simp] = linorder_class.max_less_iff_conj
   594 lemmas split_min = linorder_class.split_min
   595 lemmas split_max = linorder_class.split_max
   596 
   597 
   598 subsection {* Bounded quantifiers *}
   599 
   600 syntax
   601   "_All_less" :: "[idt, 'a, bool] => bool"    ("(3ALL _<_./ _)"  [0, 0, 10] 10)
   602   "_Ex_less" :: "[idt, 'a, bool] => bool"    ("(3EX _<_./ _)"  [0, 0, 10] 10)
   603   "_All_less_eq" :: "[idt, 'a, bool] => bool"    ("(3ALL _<=_./ _)" [0, 0, 10] 10)
   604   "_Ex_less_eq" :: "[idt, 'a, bool] => bool"    ("(3EX _<=_./ _)" [0, 0, 10] 10)
   605 
   606   "_All_greater" :: "[idt, 'a, bool] => bool"    ("(3ALL _>_./ _)"  [0, 0, 10] 10)
   607   "_Ex_greater" :: "[idt, 'a, bool] => bool"    ("(3EX _>_./ _)"  [0, 0, 10] 10)
   608   "_All_greater_eq" :: "[idt, 'a, bool] => bool"    ("(3ALL _>=_./ _)" [0, 0, 10] 10)
   609   "_Ex_greater_eq" :: "[idt, 'a, bool] => bool"    ("(3EX _>=_./ _)" [0, 0, 10] 10)
   610 
   611 syntax (xsymbols)
   612   "_All_less" :: "[idt, 'a, bool] => bool"    ("(3\<forall>_<_./ _)"  [0, 0, 10] 10)
   613   "_Ex_less" :: "[idt, 'a, bool] => bool"    ("(3\<exists>_<_./ _)"  [0, 0, 10] 10)
   614   "_All_less_eq" :: "[idt, 'a, bool] => bool"    ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
   615   "_Ex_less_eq" :: "[idt, 'a, bool] => bool"    ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
   616 
   617   "_All_greater" :: "[idt, 'a, bool] => bool"    ("(3\<forall>_>_./ _)"  [0, 0, 10] 10)
   618   "_Ex_greater" :: "[idt, 'a, bool] => bool"    ("(3\<exists>_>_./ _)"  [0, 0, 10] 10)
   619   "_All_greater_eq" :: "[idt, 'a, bool] => bool"    ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10)
   620   "_Ex_greater_eq" :: "[idt, 'a, bool] => bool"    ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10)
   621 
   622 syntax (HOL)
   623   "_All_less" :: "[idt, 'a, bool] => bool"    ("(3! _<_./ _)"  [0, 0, 10] 10)
   624   "_Ex_less" :: "[idt, 'a, bool] => bool"    ("(3? _<_./ _)"  [0, 0, 10] 10)
   625   "_All_less_eq" :: "[idt, 'a, bool] => bool"    ("(3! _<=_./ _)" [0, 0, 10] 10)
   626   "_Ex_less_eq" :: "[idt, 'a, bool] => bool"    ("(3? _<=_./ _)" [0, 0, 10] 10)
   627 
   628 syntax (HTML output)
   629   "_All_less" :: "[idt, 'a, bool] => bool"    ("(3\<forall>_<_./ _)"  [0, 0, 10] 10)
   630   "_Ex_less" :: "[idt, 'a, bool] => bool"    ("(3\<exists>_<_./ _)"  [0, 0, 10] 10)
   631   "_All_less_eq" :: "[idt, 'a, bool] => bool"    ("(3\<forall>_\<le>_./ _)" [0, 0, 10] 10)
   632   "_Ex_less_eq" :: "[idt, 'a, bool] => bool"    ("(3\<exists>_\<le>_./ _)" [0, 0, 10] 10)
   633 
   634   "_All_greater" :: "[idt, 'a, bool] => bool"    ("(3\<forall>_>_./ _)"  [0, 0, 10] 10)
   635   "_Ex_greater" :: "[idt, 'a, bool] => bool"    ("(3\<exists>_>_./ _)"  [0, 0, 10] 10)
   636   "_All_greater_eq" :: "[idt, 'a, bool] => bool"    ("(3\<forall>_\<ge>_./ _)" [0, 0, 10] 10)
   637   "_Ex_greater_eq" :: "[idt, 'a, bool] => bool"    ("(3\<exists>_\<ge>_./ _)" [0, 0, 10] 10)
   638 
   639 translations
   640   "ALL x<y. P"   =>  "ALL x. x < y \<longrightarrow> P"
   641   "EX x<y. P"    =>  "EX x. x < y \<and> P"
   642   "ALL x<=y. P"  =>  "ALL x. x <= y \<longrightarrow> P"
   643   "EX x<=y. P"   =>  "EX x. x <= y \<and> P"
   644   "ALL x>y. P"   =>  "ALL x. x > y \<longrightarrow> P"
   645   "EX x>y. P"    =>  "EX x. x > y \<and> P"
   646   "ALL x>=y. P"  =>  "ALL x. x >= y \<longrightarrow> P"
   647   "EX x>=y. P"   =>  "EX x. x >= y \<and> P"
   648 
   649 print_translation {*
   650 let
   651   val All_binder = Syntax.binder_name @{const_syntax All};
   652   val Ex_binder = Syntax.binder_name @{const_syntax Ex};
   653   val impl = @{const_syntax "op -->"};
   654   val conj = @{const_syntax "op &"};
   655   val less = @{const_syntax less};
   656   val less_eq = @{const_syntax less_eq};
   657 
   658   val trans =
   659    [((All_binder, impl, less), ("_All_less", "_All_greater")),
   660     ((All_binder, impl, less_eq), ("_All_less_eq", "_All_greater_eq")),
   661     ((Ex_binder, conj, less), ("_Ex_less", "_Ex_greater")),
   662     ((Ex_binder, conj, less_eq), ("_Ex_less_eq", "_Ex_greater_eq"))];
   663 
   664   fun matches_bound v t = 
   665      case t of (Const ("_bound", _) $ Free (v', _)) => (v = v')
   666               | _ => false
   667   fun contains_var v = Term.exists_subterm (fn Free (x, _) => x = v | _ => false)
   668   fun mk v c n P = Syntax.const c $ Syntax.mark_bound v $ n $ P
   669 
   670   fun tr' q = (q,
   671     fn [Const ("_bound", _) $ Free (v, _), Const (c, _) $ (Const (d, _) $ t $ u) $ P] =>
   672       (case AList.lookup (op =) trans (q, c, d) of
   673         NONE => raise Match
   674       | SOME (l, g) =>
   675           if matches_bound v t andalso not (contains_var v u) then mk v l u P
   676           else if matches_bound v u andalso not (contains_var v t) then mk v g t P
   677           else raise Match)
   678      | _ => raise Match);
   679 in [tr' All_binder, tr' Ex_binder] end
   680 *}
   681 
   682 
   683 subsection {* Transitivity reasoning *}
   684 
   685 lemma ord_le_eq_trans: "a <= b ==> b = c ==> a <= c"
   686 by (rule subst)
   687 
   688 lemma ord_eq_le_trans: "a = b ==> b <= c ==> a <= c"
   689 by (rule ssubst)
   690 
   691 lemma ord_less_eq_trans: "a < b ==> b = c ==> a < c"
   692 by (rule subst)
   693 
   694 lemma ord_eq_less_trans: "a = b ==> b < c ==> a < c"
   695 by (rule ssubst)
   696 
   697 lemma order_less_subst2: "(a::'a::order) < b ==> f b < (c::'c::order) ==>
   698   (!!x y. x < y ==> f x < f y) ==> f a < c"
   699 proof -
   700   assume r: "!!x y. x < y ==> f x < f y"
   701   assume "a < b" hence "f a < f b" by (rule r)
   702   also assume "f b < c"
   703   finally (order_less_trans) show ?thesis .
   704 qed
   705 
   706 lemma order_less_subst1: "(a::'a::order) < f b ==> (b::'b::order) < c ==>
   707   (!!x y. x < y ==> f x < f y) ==> a < f c"
   708 proof -
   709   assume r: "!!x y. x < y ==> f x < f y"
   710   assume "a < f b"
   711   also assume "b < c" hence "f b < f c" by (rule r)
   712   finally (order_less_trans) show ?thesis .
   713 qed
   714 
   715 lemma order_le_less_subst2: "(a::'a::order) <= b ==> f b < (c::'c::order) ==>
   716   (!!x y. x <= y ==> f x <= f y) ==> f a < c"
   717 proof -
   718   assume r: "!!x y. x <= y ==> f x <= f y"
   719   assume "a <= b" hence "f a <= f b" by (rule r)
   720   also assume "f b < c"
   721   finally (order_le_less_trans) show ?thesis .
   722 qed
   723 
   724 lemma order_le_less_subst1: "(a::'a::order) <= f b ==> (b::'b::order) < c ==>
   725   (!!x y. x < y ==> f x < f y) ==> a < f c"
   726 proof -
   727   assume r: "!!x y. x < y ==> f x < f y"
   728   assume "a <= f b"
   729   also assume "b < c" hence "f b < f c" by (rule r)
   730   finally (order_le_less_trans) show ?thesis .
   731 qed
   732 
   733 lemma order_less_le_subst2: "(a::'a::order) < b ==> f b <= (c::'c::order) ==>
   734   (!!x y. x < y ==> f x < f y) ==> f a < c"
   735 proof -
   736   assume r: "!!x y. x < y ==> f x < f y"
   737   assume "a < b" hence "f a < f b" by (rule r)
   738   also assume "f b <= c"
   739   finally (order_less_le_trans) show ?thesis .
   740 qed
   741 
   742 lemma order_less_le_subst1: "(a::'a::order) < f b ==> (b::'b::order) <= c ==>
   743   (!!x y. x <= y ==> f x <= f y) ==> a < f c"
   744 proof -
   745   assume r: "!!x y. x <= y ==> f x <= f y"
   746   assume "a < f b"
   747   also assume "b <= c" hence "f b <= f c" by (rule r)
   748   finally (order_less_le_trans) show ?thesis .
   749 qed
   750 
   751 lemma order_subst1: "(a::'a::order) <= f b ==> (b::'b::order) <= c ==>
   752   (!!x y. x <= y ==> f x <= f y) ==> a <= f c"
   753 proof -
   754   assume r: "!!x y. x <= y ==> f x <= f y"
   755   assume "a <= f b"
   756   also assume "b <= c" hence "f b <= f c" by (rule r)
   757   finally (order_trans) show ?thesis .
   758 qed
   759 
   760 lemma order_subst2: "(a::'a::order) <= b ==> f b <= (c::'c::order) ==>
   761   (!!x y. x <= y ==> f x <= f y) ==> f a <= c"
   762 proof -
   763   assume r: "!!x y. x <= y ==> f x <= f y"
   764   assume "a <= b" hence "f a <= f b" by (rule r)
   765   also assume "f b <= c"
   766   finally (order_trans) show ?thesis .
   767 qed
   768 
   769 lemma ord_le_eq_subst: "a <= b ==> f b = c ==>
   770   (!!x y. x <= y ==> f x <= f y) ==> f a <= c"
   771 proof -
   772   assume r: "!!x y. x <= y ==> f x <= f y"
   773   assume "a <= b" hence "f a <= f b" by (rule r)
   774   also assume "f b = c"
   775   finally (ord_le_eq_trans) show ?thesis .
   776 qed
   777 
   778 lemma ord_eq_le_subst: "a = f b ==> b <= c ==>
   779   (!!x y. x <= y ==> f x <= f y) ==> a <= f c"
   780 proof -
   781   assume r: "!!x y. x <= y ==> f x <= f y"
   782   assume "a = f b"
   783   also assume "b <= c" hence "f b <= f c" by (rule r)
   784   finally (ord_eq_le_trans) show ?thesis .
   785 qed
   786 
   787 lemma ord_less_eq_subst: "a < b ==> f b = c ==>
   788   (!!x y. x < y ==> f x < f y) ==> f a < c"
   789 proof -
   790   assume r: "!!x y. x < y ==> f x < f y"
   791   assume "a < b" hence "f a < f b" by (rule r)
   792   also assume "f b = c"
   793   finally (ord_less_eq_trans) show ?thesis .
   794 qed
   795 
   796 lemma ord_eq_less_subst: "a = f b ==> b < c ==>
   797   (!!x y. x < y ==> f x < f y) ==> a < f c"
   798 proof -
   799   assume r: "!!x y. x < y ==> f x < f y"
   800   assume "a = f b"
   801   also assume "b < c" hence "f b < f c" by (rule r)
   802   finally (ord_eq_less_trans) show ?thesis .
   803 qed
   804 
   805 text {*
   806   Note that this list of rules is in reverse order of priorities.
   807 *}
   808 
   809 lemmas order_trans_rules [trans] =
   810   order_less_subst2
   811   order_less_subst1
   812   order_le_less_subst2
   813   order_le_less_subst1
   814   order_less_le_subst2
   815   order_less_le_subst1
   816   order_subst2
   817   order_subst1
   818   ord_le_eq_subst
   819   ord_eq_le_subst
   820   ord_less_eq_subst
   821   ord_eq_less_subst
   822   forw_subst
   823   back_subst
   824   rev_mp
   825   mp
   826   order_neq_le_trans
   827   order_le_neq_trans
   828   order_less_trans
   829   order_less_asym'
   830   order_le_less_trans
   831   order_less_le_trans
   832   order_trans
   833   order_antisym
   834   ord_le_eq_trans
   835   ord_eq_le_trans
   836   ord_less_eq_trans
   837   ord_eq_less_trans
   838   trans
   839 
   840 
   841 (* FIXME cleanup *)
   842 
   843 text {* These support proving chains of decreasing inequalities
   844     a >= b >= c ... in Isar proofs. *}
   845 
   846 lemma xt1:
   847   "a = b ==> b > c ==> a > c"
   848   "a > b ==> b = c ==> a > c"
   849   "a = b ==> b >= c ==> a >= c"
   850   "a >= b ==> b = c ==> a >= c"
   851   "(x::'a::order) >= y ==> y >= x ==> x = y"
   852   "(x::'a::order) >= y ==> y >= z ==> x >= z"
   853   "(x::'a::order) > y ==> y >= z ==> x > z"
   854   "(x::'a::order) >= y ==> y > z ==> x > z"
   855   "(a::'a::order) > b ==> b > a ==> P"
   856   "(x::'a::order) > y ==> y > z ==> x > z"
   857   "(a::'a::order) >= b ==> a ~= b ==> a > b"
   858   "(a::'a::order) ~= b ==> a >= b ==> a > b"
   859   "a = f b ==> b > c ==> (!!x y. x > y ==> f x > f y) ==> a > f c" 
   860   "a > b ==> f b = c ==> (!!x y. x > y ==> f x > f y) ==> f a > c"
   861   "a = f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a >= f c"
   862   "a >= b ==> f b = c ==> (!! x y. x >= y ==> f x >= f y) ==> f a >= c"
   863 by auto
   864 
   865 lemma xt2:
   866   "(a::'a::order) >= f b ==> b >= c ==> (!!x y. x >= y ==> f x >= f y) ==> a >= f c"
   867 by (subgoal_tac "f b >= f c", force, force)
   868 
   869 lemma xt3: "(a::'a::order) >= b ==> (f b::'b::order) >= c ==> 
   870     (!!x y. x >= y ==> f x >= f y) ==> f a >= c"
   871 by (subgoal_tac "f a >= f b", force, force)
   872 
   873 lemma xt4: "(a::'a::order) > f b ==> (b::'b::order) >= c ==>
   874   (!!x y. x >= y ==> f x >= f y) ==> a > f c"
   875 by (subgoal_tac "f b >= f c", force, force)
   876 
   877 lemma xt5: "(a::'a::order) > b ==> (f b::'b::order) >= c==>
   878     (!!x y. x > y ==> f x > f y) ==> f a > c"
   879 by (subgoal_tac "f a > f b", force, force)
   880 
   881 lemma xt6: "(a::'a::order) >= f b ==> b > c ==>
   882     (!!x y. x > y ==> f x > f y) ==> a > f c"
   883 by (subgoal_tac "f b > f c", force, force)
   884 
   885 lemma xt7: "(a::'a::order) >= b ==> (f b::'b::order) > c ==>
   886     (!!x y. x >= y ==> f x >= f y) ==> f a > c"
   887 by (subgoal_tac "f a >= f b", force, force)
   888 
   889 lemma xt8: "(a::'a::order) > f b ==> (b::'b::order) > c ==>
   890     (!!x y. x > y ==> f x > f y) ==> a > f c"
   891 by (subgoal_tac "f b > f c", force, force)
   892 
   893 lemma xt9: "(a::'a::order) > b ==> (f b::'b::order) > c ==>
   894     (!!x y. x > y ==> f x > f y) ==> f a > c"
   895 by (subgoal_tac "f a > f b", force, force)
   896 
   897 lemmas xtrans = xt1 xt2 xt3 xt4 xt5 xt6 xt7 xt8 xt9
   898 
   899 (* 
   900   Since "a >= b" abbreviates "b <= a", the abbreviation "..." stands
   901   for the wrong thing in an Isar proof.
   902 
   903   The extra transitivity rules can be used as follows: 
   904 
   905 lemma "(a::'a::order) > z"
   906 proof -
   907   have "a >= b" (is "_ >= ?rhs")
   908     sorry
   909   also have "?rhs >= c" (is "_ >= ?rhs")
   910     sorry
   911   also (xtrans) have "?rhs = d" (is "_ = ?rhs")
   912     sorry
   913   also (xtrans) have "?rhs >= e" (is "_ >= ?rhs")
   914     sorry
   915   also (xtrans) have "?rhs > f" (is "_ > ?rhs")
   916     sorry
   917   also (xtrans) have "?rhs > z"
   918     sorry
   919   finally (xtrans) show ?thesis .
   920 qed
   921 
   922   Alternatively, one can use "declare xtrans [trans]" and then
   923   leave out the "(xtrans)" above.
   924 *)
   925 
   926 subsection {* Order on bool *}
   927 
   928 instance bool :: order 
   929   le_bool_def: "P \<le> Q \<equiv> P \<longrightarrow> Q"
   930   less_bool_def: "P < Q \<equiv> P \<le> Q \<and> P \<noteq> Q"
   931   by intro_classes (auto simp add: le_bool_def less_bool_def)
   932 lemmas [code func del] = le_bool_def less_bool_def
   933 
   934 lemma le_boolI: "(P \<Longrightarrow> Q) \<Longrightarrow> P \<le> Q"
   935 by (simp add: le_bool_def)
   936 
   937 lemma le_boolI': "P \<longrightarrow> Q \<Longrightarrow> P \<le> Q"
   938 by (simp add: le_bool_def)
   939 
   940 lemma le_boolE: "P \<le> Q \<Longrightarrow> P \<Longrightarrow> (Q \<Longrightarrow> R) \<Longrightarrow> R"
   941 by (simp add: le_bool_def)
   942 
   943 lemma le_boolD: "P \<le> Q \<Longrightarrow> P \<longrightarrow> Q"
   944 by (simp add: le_bool_def)
   945 
   946 lemma [code func]:
   947   "False \<le> b \<longleftrightarrow> True"
   948   "True \<le> b \<longleftrightarrow> b"
   949   "False < b \<longleftrightarrow> b"
   950   "True < b \<longleftrightarrow> False"
   951   unfolding le_bool_def less_bool_def by simp_all
   952 
   953 
   954 subsection {* Order on sets *}
   955 
   956 instance set :: (type) order
   957   by (intro_classes,
   958       (assumption | rule subset_refl subset_trans subset_antisym psubset_eq)+)
   959 
   960 lemmas basic_trans_rules [trans] =
   961   order_trans_rules set_rev_mp set_mp
   962 
   963 
   964 subsection {* Order on functions *}
   965 
   966 instance "fun" :: (type, ord) ord
   967   le_fun_def: "f \<le> g \<equiv> \<forall>x. f x \<le> g x"
   968   less_fun_def: "f < g \<equiv> f \<le> g \<and> f \<noteq> g" ..
   969 
   970 lemmas [code func del] = le_fun_def less_fun_def
   971 
   972 instance "fun" :: (type, order) order
   973   by default
   974     (auto simp add: le_fun_def less_fun_def expand_fun_eq
   975        intro: order_trans order_antisym)
   976 
   977 lemma le_funI: "(\<And>x. f x \<le> g x) \<Longrightarrow> f \<le> g"
   978   unfolding le_fun_def by simp
   979 
   980 lemma le_funE: "f \<le> g \<Longrightarrow> (f x \<le> g x \<Longrightarrow> P) \<Longrightarrow> P"
   981   unfolding le_fun_def by simp
   982 
   983 lemma le_funD: "f \<le> g \<Longrightarrow> f x \<le> g x"
   984   unfolding le_fun_def by simp
   985 
   986 text {*
   987   Handy introduction and elimination rules for @{text "\<le>"}
   988   on unary and binary predicates
   989 *}
   990 
   991 lemma predicate1I [Pure.intro!, intro!]:
   992   assumes PQ: "\<And>x. P x \<Longrightarrow> Q x"
   993   shows "P \<le> Q"
   994   apply (rule le_funI)
   995   apply (rule le_boolI)
   996   apply (rule PQ)
   997   apply assumption
   998   done
   999 
  1000 lemma predicate1D [Pure.dest, dest]: "P \<le> Q \<Longrightarrow> P x \<Longrightarrow> Q x"
  1001   apply (erule le_funE)
  1002   apply (erule le_boolE)
  1003   apply assumption+
  1004   done
  1005 
  1006 lemma predicate2I [Pure.intro!, intro!]:
  1007   assumes PQ: "\<And>x y. P x y \<Longrightarrow> Q x y"
  1008   shows "P \<le> Q"
  1009   apply (rule le_funI)+
  1010   apply (rule le_boolI)
  1011   apply (rule PQ)
  1012   apply assumption
  1013   done
  1014 
  1015 lemma predicate2D [Pure.dest, dest]: "P \<le> Q \<Longrightarrow> P x y \<Longrightarrow> Q x y"
  1016   apply (erule le_funE)+
  1017   apply (erule le_boolE)
  1018   apply assumption+
  1019   done
  1020 
  1021 lemma rev_predicate1D: "P x ==> P <= Q ==> Q x"
  1022   by (rule predicate1D)
  1023 
  1024 lemma rev_predicate2D: "P x y ==> P <= Q ==> Q x y"
  1025   by (rule predicate2D)
  1026 
  1027 
  1028 subsection {* Monotonicity, least value operator and min/max *}
  1029 
  1030 locale mono =
  1031   fixes f
  1032   assumes mono: "A \<le> B \<Longrightarrow> f A \<le> f B"
  1033 
  1034 lemmas monoI [intro?] = mono.intro
  1035   and monoD [dest?] = mono.mono
  1036 
  1037 lemma LeastI2_order:
  1038   "[| P (x::'a::order);
  1039       !!y. P y ==> x <= y;
  1040       !!x. [| P x; ALL y. P y --> x \<le> y |] ==> Q x |]
  1041    ==> Q (Least P)"
  1042 apply (unfold Least_def)
  1043 apply (rule theI2)
  1044   apply (blast intro: order_antisym)+
  1045 done
  1046 
  1047 lemma Least_mono:
  1048   "mono (f::'a::order => 'b::order) ==> EX x:S. ALL y:S. x <= y
  1049     ==> (LEAST y. y : f ` S) = f (LEAST x. x : S)"
  1050     -- {* Courtesy of Stephan Merz *}
  1051   apply clarify
  1052   apply (erule_tac P = "%x. x : S" in LeastI2_order, fast)
  1053   apply (rule LeastI2_order)
  1054   apply (auto elim: monoD intro!: order_antisym)
  1055   done
  1056 
  1057 lemma Least_equality:
  1058   "[| P (k::'a::order); !!x. P x ==> k <= x |] ==> (LEAST x. P x) = k"
  1059 apply (simp add: Least_def)
  1060 apply (rule the_equality)
  1061 apply (auto intro!: order_antisym)
  1062 done
  1063 
  1064 lemma min_leastL: "(!!x. least <= x) ==> min least x = least"
  1065 by (simp add: min_def)
  1066 
  1067 lemma max_leastL: "(!!x. least <= x) ==> max least x = x"
  1068 by (simp add: max_def)
  1069 
  1070 lemma min_leastR: "(\<And>x\<Colon>'a\<Colon>order. least \<le> x) \<Longrightarrow> min x least = least"
  1071 apply (simp add: min_def)
  1072 apply (blast intro: order_antisym)
  1073 done
  1074 
  1075 lemma max_leastR: "(\<And>x\<Colon>'a\<Colon>order. least \<le> x) \<Longrightarrow> max x least = x"
  1076 apply (simp add: max_def)
  1077 apply (blast intro: order_antisym)
  1078 done
  1079 
  1080 lemma min_of_mono:
  1081   "(!!x y. (f x <= f y) = (x <= y)) ==> min (f m) (f n) = f (min m n)"
  1082 by (simp add: min_def)
  1083 
  1084 lemma max_of_mono:
  1085   "(!!x y. (f x <= f y) = (x <= y)) ==> max (f m) (f n) = f (max m n)"
  1086 by (simp add: max_def)
  1087 
  1088 
  1089 subsection {* legacy ML bindings *}
  1090 
  1091 ML {*
  1092 val monoI = @{thm monoI};
  1093 *}
  1094 
  1095 end