author | wenzelm |
Sat, 07 Sep 2019 16:17:30 +0200 | |
changeset 70672 | e4bba654d085 |
parent 69597 | ff784d5a5bfb |
permissions | -rw-r--r-- |
27063 | 1 |
theory Examples2 |
2 |
imports Examples |
|
3 |
begin |
|
67399 | 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 | 6 |
proof - |
61419 | 7 |
txt \<open>\normalsize The goals are now: |
32983 | 8 |
@{subgoals [display]} |
69505 | 9 |
The proof that~\<open>\<le>\<close> is a partial order is as above.\<close> |
67399 | 10 |
show "partial_order ((\<le>) :: int \<Rightarrow> int \<Rightarrow> bool)" |
32981 | 11 |
by unfold_locales auto |
61419 | 12 |
txt \<open>\normalsize The second goal is shown by unfolding the |
69597 | 13 |
definition of \<^term>\<open>partial_order.less\<close>.\<close> |
67399 | 14 |
show "partial_order.less (\<le>) x y = (x < y)" |
15 |
unfolding partial_order.less_def [OF \<open>partial_order (\<le>)\<close>] |
|
32981 | 16 |
by auto |
17 |
qed |
|
27063 | 18 |
|
61419 | 19 |
text \<open>Note that the above proof is not in the context of the |
69505 | 20 |
interpreted locale. Hence, the premise of \<open>partial_order.less_def\<close> is discharged manually with \<open>OF\<close>. |
61419 | 21 |
\<close> |
27063 | 22 |
end |