X-Git-Url: http://git.megacz.com/?p=coq-hetmet.git;a=blobdiff_plain;f=src%2FHaskProof.v;h=efb9220c75118dec6e7237beb9efd9f5242ad0d4;hp=7ae3462877aafbb9115538ab0fb9f5ecedd29d2b;hb=525c4131ed6329017b4738f0b64a1d8f6c1eb164;hpb=53d4f1ce851b924cab5dc39419179a366001cbca diff --git a/src/HaskProof.v b/src/HaskProof.v index 7ae3462..efb9220 100644 --- a/src/HaskProof.v +++ b/src/HaskProof.v @@ -40,8 +40,8 @@ Inductive UJudg (Γ:TypeEnv)(Δ:CoercionEnv Γ) := Tree ??(LeveledHaskType Γ ★) -> UJudg Γ Δ. Notation "Γ >> Δ > a '|-' s" := (mkUJudg Γ Δ a s) (at level 52, Δ at level 50, a at level 52, s at level 50). - Notation "'ext_tree_left'" := (fun ctx j => match j with mkUJudg t s => mkUJudg _ _ (ctx,,t) s end). - Notation "'ext_tree_right'" := (fun ctx j => match j with mkUJudg t s => mkUJudg _ _ (t,,ctx) s end). + Definition ext_tree_left {Γ}{Δ} := (fun ctx (j:UJudg Γ Δ) => match j with mkUJudg t s => mkUJudg Γ Δ (ctx,,t) s end). + Definition ext_tree_right {Γ}{Δ} := (fun ctx (j:UJudg Γ Δ) => match j with mkUJudg t s => mkUJudg Γ Δ (t,,ctx) s end). (* we can turn a UJudg into a Judg by simply internalizing the index *) Definition UJudg2judg {Γ}{Δ}(ej:@UJudg Γ Δ) : Judg :=