doc-src/Locales/Examples2.thy
author wenzelm
Tue, 28 Aug 2012 16:43:47 +0200
changeset 48971 5a4bcf466156
parent 48943 54da920baf38
permissions -rw-r--r--
prepare document more uniformly; fewer latex runs, in accordance to "isabelle document"; apply fixbookmarks last to operate on final root.out;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27063
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     1
theory Examples2
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     2
imports Examples
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
     3
begin
32981
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
     4
text {* \vspace{-5ex} *}
32982
40810d98f4c9 Changed part of the examples to int.
ballarin
parents: 32981
diff changeset
     5
  interpretation %visible int: partial_order "op \<le> :: [int, int] \<Rightarrow> bool"
43543
eb8b4851b039 While reading equations of an interpretation, already allow syntax provided by the interpretation base.
ballarin
parents: 32983
diff changeset
     6
    where "int.less x y = (x < y)"
32981
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
     7
  proof -
32983
a6914429005b Finished revisions of locales tutorial.
ballarin
parents: 32982
diff changeset
     8
    txt {* \normalsize The goals are now:
a6914429005b Finished revisions of locales tutorial.
ballarin
parents: 32982
diff changeset
     9
      @{subgoals [display]}
a6914429005b Finished revisions of locales tutorial.
ballarin
parents: 32982
diff changeset
    10
      The proof that~@{text \<le>} is a partial order is as above. *}
32982
40810d98f4c9 Changed part of the examples to int.
ballarin
parents: 32981
diff changeset
    11
    show "partial_order (op \<le> :: int \<Rightarrow> int \<Rightarrow> bool)"
32981
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    12
      by unfold_locales auto
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    13
    txt {* \normalsize The second goal is shown by unfolding the
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    14
      definition of @{term "partial_order.less"}. *}
43543
eb8b4851b039 While reading equations of an interpretation, already allow syntax provided by the interpretation base.
ballarin
parents: 32983
diff changeset
    15
    show "partial_order.less op \<le> x y = (x < y)"
32981
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    16
      unfolding partial_order.less_def [OF `partial_order op \<le>`]
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    17
      by auto
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    18
  qed
27063
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    19
32981
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    20
text {* Note that the above proof is not in the context of the
32983
a6914429005b Finished revisions of locales tutorial.
ballarin
parents: 32982
diff changeset
    21
  interpreted locale.  Hence, the premise of @{text
a6914429005b Finished revisions of locales tutorial.
ballarin
parents: 32982
diff changeset
    22
  "partial_order.less_def"} is discharged manually with @{text OF}.
27063
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    23
  *}
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    24
end