src/Doc/Locales/Examples2.thy
author nipkow
Sun, 06 May 2018 13:51:37 +0200
changeset 68083 d730a8cfc6e0
parent 67399 eab6ce8368fa
child 69505 cc2d676d5395
permissions -rw-r--r--
removed asm "finite"
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
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67349
diff changeset
     4
  interpretation %visible int: partial_order "(\<le>) :: [int, int] \<Rightarrow> bool"
61566
c3d6e570ccef Keyword 'rewrites' identifies rewrite morphisms.
ballarin
parents: 61419
diff changeset
     5
    rewrites "int.less x y = (x < y)"
32981
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
     6
  proof -
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 57607
diff changeset
     7
    txt \<open>\normalsize The goals are now:
32983
a6914429005b Finished revisions of locales tutorial.
ballarin
parents: 32982
diff changeset
     8
      @{subgoals [display]}
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 57607
diff changeset
     9
      The proof that~@{text \<le>} is a partial order is as above.\<close>
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67349
diff changeset
    10
    show "partial_order ((\<le>) :: int \<Rightarrow> int \<Rightarrow> bool)"
32981
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    11
      by unfold_locales auto
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 57607
diff changeset
    12
    txt \<open>\normalsize The second goal is shown by unfolding the
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 57607
diff changeset
    13
      definition of @{term "partial_order.less"}.\<close>
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67349
diff changeset
    14
    show "partial_order.less (\<le>) x y = (x < y)"
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67349
diff changeset
    15
      unfolding partial_order.less_def [OF \<open>partial_order (\<le>)\<close>]
32981
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    16
      by auto
0114e04a0d64 Save current state of locales tutorial.
ballarin
parents: 30826
diff changeset
    17
  qed
27063
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    18
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 57607
diff changeset
    19
text \<open>Note that the above proof is not in the context of the
32983
a6914429005b Finished revisions of locales tutorial.
ballarin
parents: 32982
diff changeset
    20
  interpreted locale.  Hence, the premise of @{text
a6914429005b Finished revisions of locales tutorial.
ballarin
parents: 32982
diff changeset
    21
  "partial_order.less_def"} is discharged manually with @{text OF}.
61419
3c3f8b182e4b isabelle update_cartouches;
wenzelm
parents: 57607
diff changeset
    22
\<close>
27063
d1d35284542f New version covering interpretation.
ballarin
parents:
diff changeset
    23
end