prio/Precedence_ord.thy
author urbanc
Mon, 30 Jan 2012 08:55:27 +0000
changeset 269 70395e3fd99f
parent 262 4190df6f4488
child 320 630754a81bdb
permissions -rw-r--r--
more text
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
262
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     1
(*  Title:      HOL/Library/Product_ord.thy
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     2
    Author:     Norbert Voelker
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     3
*)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     4
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     5
header {* Order on product types *}
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     6
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     7
theory Precedence_ord
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     8
imports Main
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
     9
begin
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    10
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    11
datatype precedence = Prc nat nat
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    12
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    13
instantiation precedence :: order
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    14
begin
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    15
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    16
definition
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    17
  precedence_le_def: "x \<le> y \<longleftrightarrow> (case (x, y) of
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    18
                                   (Prc fx sx, Prc fy sy) \<Rightarrow> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    19
                                 fx < fy \<or> (fx \<le> fy \<and> sy \<le> sx))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    20
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    21
definition
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    22
  precedence_less_def: "x < y \<longleftrightarrow> (case (x, y) of
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    23
                               (Prc fx sx, Prc fy sy) \<Rightarrow> 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    24
                                     fx < fy \<or> (fx \<le> fy \<and> sy < sx))"
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    25
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    26
instance
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    27
proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    28
qed (auto simp: precedence_le_def precedence_less_def 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    29
              intro: order_trans split:precedence.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    30
end
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    31
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    32
instance precedence :: preorder ..
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    33
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    34
instance precedence :: linorder proof
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    35
qed (auto simp: precedence_le_def precedence_less_def 
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    36
              intro: order_trans split:precedence.splits)
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    37
4190df6f4488 initial version of the PIP formalisation
urbanc
parents:
diff changeset
    38
end