Nominal/Ex/Datatypes.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 10 Apr 2012 16:02:30 +0100
changeset 3158 89f9d7e85e88
parent 3065 51ef8a3cb6ef
child 3208 da575186d492
child 3245 017e33849f4d
permissions -rw-r--r--
moved lift_raw_const from Quotient to Nominal
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2559
add799cf0817 adapted to changes by Florian on the quotient package and removed local fix for function package
Christian Urban <urbanc@in.tum.de>
parents: 2557
diff changeset
     1
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
theory Datatypes
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
imports "../Nominal2" 
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
begin
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
(* 
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
  plain nominal_datatype definition without an 
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  atom_decl - example by John Matthews
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
*)
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
3065
51ef8a3cb6ef updated to lates changes in the datatype package
Christian Urban <urbanc@in.tum.de>
parents: 3046
diff changeset
    11
nominal_datatype 'a::fs Maybe = 
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
  Nothing 
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
| Just 'a
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
thm Maybe.distinct
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
thm Maybe.induct
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
thm Maybe.exhaust
2617
e44551d067e6 properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents: 2559
diff changeset
    19
thm Maybe.strong_exhaust
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
thm Maybe.fv_defs
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
thm Maybe.bn_defs
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
thm Maybe.perm_simps
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
thm Maybe.eq_iff
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
thm Maybe.fv_bn_eqvt
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
thm Maybe.size_eqvt
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
thm Maybe.supp
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
(*
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
  using a datatype inside a nominal_datatype
2617
e44551d067e6 properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents: 2559
diff changeset
    30
e44551d067e6 properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents: 2559
diff changeset
    31
  the datatype needs to be shown to be an instance
e44551d067e6 properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents: 2559
diff changeset
    32
  of the pure class (this is usually trivial)
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
*)
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
atom_decl name
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
datatype foo = Foo | Bar
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
instantiation foo :: pure
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
begin
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
definition
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  "p \<bullet> (f::foo) = f"
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
instance
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
apply(default)
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
apply(simp_all add:  permute_foo_def)
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
done
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
end
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
nominal_datatype baz =
2950
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2868
diff changeset
    55
  Baz x::name t::foo   binds x in t
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
2868
2b8e387d2dfc got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de>
parents: 2781
diff changeset
    58
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
thm baz.distinct
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
thm baz.induct
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
thm baz.exhaust
2617
e44551d067e6 properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents: 2559
diff changeset
    62
thm baz.strong_exhaust
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
thm baz.fv_defs
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
thm baz.bn_defs
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
thm baz.perm_simps
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
thm baz.eq_iff
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
thm baz.fv_bn_eqvt
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
thm baz.size_eqvt
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
thm baz.supp
2777
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    70
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    71
(*
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    72
  a nominal datatype with a set argument of
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    73
  pure type
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    74
*)
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
  
2777
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    76
nominal_datatype set_ty = 
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    77
  Set_ty "nat set"
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    78
| Fun "nat \<Rightarrow> nat"
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    79
| Var "name"
2950
0911cb7bf696 changed bind to binds in specifications; bind will cause trouble with Monad_Syntax
Christian Urban <urbanc@in.tum.de>
parents: 2868
diff changeset
    80
| Lam x::"name" t::"set_ty" binds x in t
2868
2b8e387d2dfc got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de>
parents: 2781
diff changeset
    81
thm meta_eq_to_obj_eq
2777
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    82
thm set_ty.distinct
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    83
thm set_ty.induct
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    84
thm set_ty.exhaust
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    85
thm set_ty.strong_exhaust
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    86
thm set_ty.fv_defs
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    87
thm set_ty.bn_defs
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    88
thm set_ty.perm_simps
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    89
thm set_ty.eq_iff
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    90
thm set_ty.fv_bn_eqvt
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    91
thm set_ty.size_eqvt
75a95431cd8b proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents: 2617
diff changeset
    92
thm set_ty.supp
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
2781
542ff50555f5 updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents: 2777
diff changeset
    94
542ff50555f5 updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents: 2777
diff changeset
    95
2485
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
end
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
6bab47906dbe added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99