author | paulson <lp15@cam.ac.uk> |
Tue, 31 Mar 2015 15:01:06 +0100 | |
changeset 59863 | 30519ff3dffb |
parent 57607 | 5ff0cf3f5f6f |
child 61419 | 3c3f8b182e4b |
permissions | -rw-r--r-- |
27063 | 1 |
theory Examples2 |
2 |
imports Examples |
|
3 |
begin |
|
32982 | 4 |
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
|
5 |
where "int.less x y = (x < y)" |
32981 | 6 |
proof - |
32983 | 7 |
txt {* \normalsize The goals are now: |
8 |
@{subgoals [display]} |
|
9 |
The proof that~@{text \<le>} is a partial order is as above. *} |
|
32982 | 10 |
show "partial_order (op \<le> :: int \<Rightarrow> int \<Rightarrow> bool)" |
32981 | 11 |
by unfold_locales auto |
12 |
txt {* \normalsize The second goal is shown by unfolding the |
|
13 |
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
|
14 |
show "partial_order.less op \<le> x y = (x < y)" |
32981 | 15 |
unfolding partial_order.less_def [OF `partial_order op \<le>`] |
16 |
by auto |
|
17 |
qed |
|
27063 | 18 |
|
32981 | 19 |
text {* Note that the above proof is not in the context of the |
32983 | 20 |
interpreted locale. Hence, the premise of @{text |
21 |
"partial_order.less_def"} is discharged manually with @{text OF}. |
|
27063 | 22 |
*} |
23 |
end |