doc-src/Locales/Locales/Examples2.thy
author immler@in.tum.de
Thu, 26 Feb 2009 10:13:43 +0100
changeset 30151 629f3a92863e
parent 29293 d4ef21262b8f
child 30580 cc5a55d7a5be
permissions -rw-r--r--
removed global ref dfg_format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27081
6d2a458be1b6 replaced (*<*)(*>*) by invisibility tags;
wenzelm
parents: 27063
diff changeset
     1
(* $Id$ *)
6d2a458be1b6 replaced (*<*)(*>*) by invisibility tags;
wenzelm
parents: 27063
diff changeset
     2
27063
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     3
theory Examples2
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     4
imports Examples
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     5
begin
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     6
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     7
text {* This is achieved by unfolding suitable equations during
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     8
  interpretation.  These equations are given after the keyword
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     9
  \isakeyword{where} and require proofs.  The revised command,
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    10
  replacing @{text "\<sqsubset>"} by @{text "<"}, is: *}
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    11
29293
d4ef21262b8f crude adaption to new locales;
wenzelm
parents: 27081
diff changeset
    12
interpretation %visible nat: partial_order "op \<le> :: [nat, nat] \<Rightarrow> bool"
27063
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    13
  where "partial_order.less op \<le> (x::nat) y = (x < y)"
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    14
proof -
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    15
  txt {* The goals are @{subgoals [display]}
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    16
    The proof that @{text \<le>} is a partial order is a above. *}
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    17
  show "partial_order (op \<le> :: nat \<Rightarrow> nat \<Rightarrow> bool)"
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    18
    by unfold_locales auto
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    19
  txt {* The second goal is shown by unfolding the
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    20
    definition of @{term "partial_order.less"}. *}
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    21
  show "partial_order.less op \<le> (x::nat) y = (x < y)"
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    22
    unfolding partial_order.less_def [OF `partial_order op \<le>`]
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    23
    by auto
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    24
qed
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    25
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    26
text {* Note that the above proof is not in the context of a locale.
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    27
  Hence, the correct interpretation of @{text
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    28
  "partial_order.less_def"} is obtained manually with @{text OF}.
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    29
  *}
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    30
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    31
end