src/HOL/Quotient_Examples/Lift_Set.thy
author bulwahn
Thu, 17 Nov 2011 19:01:05 +0100
changeset 45536 5b0b1dc2e40f
child 45576 6ea2bba2694a
permissions -rw-r--r--
adding a preliminary example to show how the quotient_definition package can be generalized
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
45536
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
     1
(*  Title:      HOL/Quotient.thy
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
     2
    Author:     Lukas Bulwahn and Ondrey Kuncar
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
     3
*)
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
     4
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
     5
header {* Example of lifting definitions with the quotient infrastructure *}
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
     6
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
     7
theory Lift_Set
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
     8
imports Main
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
     9
begin
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    10
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    11
typedef 'a set = "(UNIV :: ('a => bool) => bool)"
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    12
morphisms member Set by auto
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    13
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    14
text {* Here is some ML setup that should eventually be incorporated in the typedef command. *}
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    15
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    16
local_setup {* fn lthy =>
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    17
let
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    18
  val quotients = {qtyp = @{typ "'a set"}, rtyp = @{typ "'a => bool"}, equiv_rel = @{term "dummy"}, equiv_thm = @{thm refl}}
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    19
  val qty_full_name = @{type_name "set"}
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    20
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    21
  fun qinfo phi = Quotient_Info.transform_quotients phi quotients
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    22
  in lthy
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    23
    |> Local_Theory.declaration {syntax = false, pervasive = true}
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    24
        (fn phi => Quotient_Info.update_quotients qty_full_name (qinfo phi)
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    25
       #> Quotient_Info.update_abs_rep qty_full_name (Quotient_Info.transform_abs_rep phi {abs = @{term "Set"}, rep = @{term "member"}}))
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    26
  end
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    27
*}
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    28
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    29
text {* Now, we can employ quotient_definition to lift definitions. *}
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    30
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    31
quotient_definition empty where "empty :: 'a set"
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    32
is "Set.empty"
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    33
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    34
term "Lift_Set.empty"
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    35
thm Lift_Set.empty_def
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    36
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    37
quotient_definition insert where "insert :: 'a => 'a set => 'a set"
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    38
is "Set.insert"
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    39
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    40
term "Lift_Set.insert"
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    41
thm Lift_Set.insert_def
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    42
5b0b1dc2e40f adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff changeset
    43
end